--- a/NEWS Thu Jun 18 17:14:08 2009 +0100
+++ b/NEWS Tue Jun 23 15:48:55 2009 +0100
@@ -43,6 +43,18 @@
* Constants Set.Pow and Set.image now with authentic syntax; object-logic definitions
Set.Pow_def and Set.image_def. INCOMPATIBILITY.
+* Discontinued ancient tradition to suffix certain ML module names with "_package", e.g.:
+
+ DatatypePackage ~> Datatype
+ InductivePackage ~> Inductive
+
+ etc.
+
+INCOMPATIBILITY.
+
+
+* NewNumberTheory: Jeremy Avigad's new version of part of NumberTheory.
+If possible, use NewNumberTheory, not NumberTheory.
*** ML ***
--- a/doc-src/Locales/Locales/Examples3.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/doc-src/Locales/Locales/Examples3.thy Tue Jun 23 15:48:55 2009 +0100
@@ -143,8 +143,8 @@
interpretation nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
where "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
- and nat_dvd_meet_eq: "lattice.meet op dvd = gcd"
- and nat_dvd_join_eq: "lattice.join op dvd = lcm"
+ and nat_dvd_meet_eq: "lattice.meet (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = gcd"
+ and nat_dvd_join_eq: "lattice.join (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = lcm"
proof -
show "lattice (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool)"
apply unfold_locales
@@ -152,23 +152,24 @@
apply (rule_tac x = "gcd x y" in exI)
apply auto [1]
apply (rule_tac x = "lcm x y" in exI)
- apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
+ apply (auto intro: nat_lcm_least)
done
then interpret nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
show "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
by (rule nat_dvd_less_eq)
- show "lattice.meet op dvd = gcd"
+ show "lattice.meet (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = gcd"
apply (auto simp add: expand_fun_eq)
apply (unfold nat_dvd.meet_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_inf_def)
by auto
- show "lattice.join op dvd = lcm"
+ show "lattice.join (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = lcm"
apply (auto simp add: expand_fun_eq)
apply (unfold nat_dvd.join_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_sup_def)
- by (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
+ apply (auto intro: nat_lcm_least iff: nat_lcm_unique)
+ done
qed
text {* Equations @{thm [source] nat_dvd_meet_eq} and @{thm [source]
@@ -198,8 +199,8 @@
interpretation %visible nat_dvd:
distrib_lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
where "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
- and "lattice.meet op dvd = gcd"
- and "lattice.join op dvd = lcm"
+ and "lattice.meet (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = gcd"
+ and "lattice.join (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = lcm"
proof -
show "distrib_lattice (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool)"
apply unfold_locales
--- a/doc-src/Locales/Locales/document/Examples3.tex Thu Jun 18 17:14:08 2009 +0100
+++ b/doc-src/Locales/Locales/document/Examples3.tex Tue Jun 23 15:48:55 2009 +0100
@@ -307,8 +307,8 @@
\isacommand{interpretation}\isamarkupfalse%
\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ dvd\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
@@ -330,7 +330,7 @@
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}lcm\ x\ y{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}dvd{\isadigit{1}}\ lcm{\isacharunderscore}dvd{\isadigit{2}}\ lcm{\isacharunderscore}least{\isacharparenright}\isanewline
+\ {\isacharparenleft}auto\ intro{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}least{\isacharparenright}\isanewline
\ \ \ \ \isacommand{done}\isamarkupfalse%
\isanewline
\ \ \isacommand{then}\isamarkupfalse%
@@ -342,7 +342,7 @@
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq{\isacharparenright}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
+\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
@@ -354,7 +354,7 @@
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ dvd\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
+\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
@@ -363,8 +363,10 @@
\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}dvd{\isadigit{1}}\ lcm{\isacharunderscore}dvd{\isadigit{2}}\ lcm{\isacharunderscore}least{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}auto\ intro{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}least\ iff{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}unique{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{done}\isamarkupfalse%
+\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
@@ -407,8 +409,8 @@
\ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline
\ \ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ dvd\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
--- a/src/HOL/Algebra/Bij.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/Bij.thy Tue Jun 23 15:48:55 2009 +0100
@@ -50,7 +50,7 @@
apply (simp add: compose_Bij)
apply (simp add: id_Bij)
apply (simp add: compose_Bij)
- apply (blast intro: compose_assoc [symmetric] Bij_imp_funcset)
+ apply (blast intro: compose_assoc [symmetric] dest: Bij_imp_funcset)
apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp)
apply (blast intro: Bij_compose_restrict_eq restrict_Inv_Bij)
done
--- a/src/HOL/Algebra/Coset.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/Coset.thy Tue Jun 23 15:48:55 2009 +0100
@@ -934,8 +934,8 @@
lemma (in group_hom) FactGroup_hom:
"(\<lambda>X. contents (h`X)) \<in> hom (G Mod (kernel G H h)) H"
-apply (simp add: hom_def FactGroup_contents_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed)
-proof (simp add: hom_def funcsetI FactGroup_contents_mem, intro ballI)
+apply (simp add: hom_def FactGroup_contents_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed)
+proof (intro ballI)
fix X and X'
assume X: "X \<in> carrier (G Mod kernel G H h)"
and X': "X' \<in> carrier (G Mod kernel G H h)"
@@ -952,7 +952,7 @@
simp add: set_mult_def image_eq_UN
subsetD [OF Xsub] subsetD [OF X'sub])
thus "contents (h ` (X <#> X')) = contents (h ` X) \<otimes>\<^bsub>H\<^esub> contents (h ` X')"
- by (simp add: all image_eq_UN FactGroup_nonempty X X')
+ by (simp add: all image_eq_UN FactGroup_nonempty X X')
qed
--- a/src/HOL/Algebra/Exponent.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/Exponent.thy Tue Jun 23 15:48:55 2009 +0100
@@ -9,6 +9,8 @@
imports Main Primes Binomial
begin
+hide (open) const GCD.gcd GCD.coprime GCD.prime
+
section {*Sylow's Theorem*}
subsection {*The Combinatorial Argument Underlying the First Sylow Theorem*}
--- a/src/HOL/Algebra/FiniteProduct.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/FiniteProduct.thy Tue Jun 23 15:48:55 2009 +0100
@@ -508,7 +508,6 @@
apply force
apply (subst finprod_insert)
apply auto
- apply (force simp add: Pi_def)
apply (subst m_comm)
apply auto
done
--- a/src/HOL/Algebra/Group.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/Group.thy Tue Jun 23 15:48:55 2009 +0100
@@ -542,24 +542,22 @@
(\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y)}"
lemma (in group) hom_compose:
- "[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I"
-apply (auto simp add: hom_def funcset_compose)
-apply (simp add: compose_def funcset_mem)
-done
+ "[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I"
+by (fastsimp simp add: hom_def compose_def)
constdefs
iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60)
"G \<cong> H == {h. h \<in> hom G H & bij_betw h (carrier G) (carrier H)}"
lemma iso_refl: "(%x. x) \<in> G \<cong> G"
-by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
+by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
lemma (in group) iso_sym:
"h \<in> G \<cong> H \<Longrightarrow> Inv (carrier G) h \<in> H \<cong> G"
apply (simp add: iso_def bij_betw_Inv)
apply (subgoal_tac "Inv (carrier G) h \<in> carrier H \<rightarrow> carrier G")
prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_Inv])
-apply (simp add: hom_def bij_betw_def Inv_f_eq funcset_mem f_Inv_f)
+apply (simp add: hom_def bij_betw_def Inv_f_eq f_Inv_f Pi_def)
done
lemma (in group) iso_trans:
@@ -568,11 +566,11 @@
lemma DirProd_commute_iso:
shows "(\<lambda>(x,y). (y,x)) \<in> (G \<times>\<times> H) \<cong> (H \<times>\<times> G)"
-by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
+by (auto simp add: iso_def hom_def inj_on_def bij_betw_def)
lemma DirProd_assoc_iso:
shows "(\<lambda>(x,y,z). (x,(y,z))) \<in> (G \<times>\<times> H \<times>\<times> I) \<cong> (G \<times>\<times> (H \<times>\<times> I))"
-by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
+by (auto simp add: iso_def hom_def inj_on_def bij_betw_def)
text{*Basis for homomorphism proofs: we assume two groups @{term G} and
@@ -592,7 +590,7 @@
"x \<in> carrier G ==> h x \<in> carrier H"
proof -
assume "x \<in> carrier G"
- with homh [unfolded hom_def] show ?thesis by (auto simp add: funcset_mem)
+ with homh [unfolded hom_def] show ?thesis by auto
qed
lemma (in group_hom) one_closed [simp]:
--- a/src/HOL/Algebra/Sylow.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Algebra/Sylow.thy Tue Jun 23 15:48:55 2009 +0100
@@ -371,4 +371,3 @@
done
end
-
--- a/src/HOL/Code_Eval.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Code_Eval.thy Tue Jun 23 15:48:55 2009 +0100
@@ -85,7 +85,9 @@
end;
fun add_term_of_code tyco raw_vs raw_cs thy =
let
- val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
+ val algebra = Sign.classes_of thy;
+ val vs = map (fn (v, sort) =>
+ (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
val ty = Type (tyco, map TFree vs);
val cs = (map o apsnd o map o map_atyps)
(fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
--- a/src/HOL/Decision_Procs/Cooper.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1071,8 +1071,8 @@
"plusinf p = p"
recdef \<delta> "measure size"
- "\<delta> (And p q) = zlcm (\<delta> p) (\<delta> q)"
- "\<delta> (Or p q) = zlcm (\<delta> p) (\<delta> q)"
+ "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
+ "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
"\<delta> (Dvd i (CN 0 c e)) = i"
"\<delta> (NDvd i (CN 0 c e)) = i"
"\<delta> p = 1"
@@ -1104,20 +1104,20 @@
proof (induct p rule: iszlfm.induct)
case (1 p q)
let ?d = "\<delta> (And p q)"
- from prems zlcm_pos have dp: "?d >0" by simp
+ from prems int_lcm_pos have dp: "?d >0" by simp
have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp
- hence th: "d\<delta> p ?d" using delta_mono prems(3-4) by(simp del:dvd_zlcm_self1)
+ hence th: "d\<delta> p ?d" using delta_mono prems(3-4) by(simp only: iszlfm.simps)
have "\<delta> q dvd \<delta> (And p q)" using prems by simp
- hence th': "d\<delta> q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2)
+ hence th': "d\<delta> q ?d" using delta_mono prems by(simp only: iszlfm.simps)
from th th' dp show ?case by simp
next
case (2 p q)
let ?d = "\<delta> (And p q)"
- from prems zlcm_pos have dp: "?d >0" by simp
+ from prems int_lcm_pos have dp: "?d >0" by simp
have "\<delta> p dvd \<delta> (And p q)" using prems by simp
- hence th: "d\<delta> p ?d" using delta_mono prems by(simp del:dvd_zlcm_self1)
+ hence th: "d\<delta> p ?d" using delta_mono prems by(simp only: iszlfm.simps)
have "\<delta> q dvd \<delta> (And p q)" using prems by simp
- hence th': "d\<delta> q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2)
+ hence th': "d\<delta> q ?d" using delta_mono prems by(simp only: iszlfm.simps)
from th th' dp show ?case by simp
qed simp_all
@@ -1156,8 +1156,8 @@
"d\<beta> p = (\<lambda> k. True)"
recdef \<zeta> "measure size"
- "\<zeta> (And p q) = zlcm (\<zeta> p) (\<zeta> q)"
- "\<zeta> (Or p q) = zlcm (\<zeta> p) (\<zeta> q)"
+ "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
+ "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
"\<zeta> (Eq (CN 0 c e)) = c"
"\<zeta> (NEq (CN 0 c e)) = c"
"\<zeta> (Lt (CN 0 c e)) = c"
@@ -1406,19 +1406,19 @@
using linp
proof(induct p rule: iszlfm.induct)
case (1 p q)
- from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: zlcm_pos)
+ from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
next
case (2 p q)
- from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: zlcm_pos)
-qed (auto simp add: zlcm_pos)
+ from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
+qed (auto simp add: int_lcm_pos)
lemma a\<beta>: assumes linp: "iszlfm p" and d: "d\<beta> p l" and lp: "l > 0"
shows "iszlfm (a\<beta> p l) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm bbs (l*x #bs) (a\<beta> p l) = Ifm bbs (x#bs) p)"
--- a/src/HOL/Decision_Procs/Ferrack.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy Tue Jun 23 15:48:55 2009 +0100
@@ -478,8 +478,8 @@
by (induct t rule: maxcoeff.induct, auto)
recdef numgcdh "measure size"
- "numgcdh (C i) = (\<lambda>g. zgcd i g)"
- "numgcdh (CN n c t) = (\<lambda>g. zgcd c (numgcdh t g))"
+ "numgcdh (C i) = (\<lambda>g. gcd i g)"
+ "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
"numgcdh t = (\<lambda>g. 1)"
defs numgcd_def [code]: "numgcd t \<equiv> numgcdh t (maxcoeff t)"
@@ -512,11 +512,11 @@
assumes g0: "numgcd t = 0"
shows "Inum bs t = 0"
using g0[simplified numgcd_def]
- by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos)
+ by (induct t rule: numgcdh.induct, auto simp add: natabs0 max_def maxcoeff_pos)
lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
using gp
- by (induct t rule: numgcdh.induct, auto simp add: zgcd_def)
+ by (induct t rule: numgcdh.induct, auto)
lemma numgcd_pos: "numgcd t \<ge>0"
by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
@@ -551,15 +551,15 @@
from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1)
qed simp_all
-lemma zgcd_gt1: "zgcd i j > 1 \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
- apply (cases "abs i = 0", simp_all add: zgcd_def)
+lemma zgcd_gt1: "gcd i j > (1::int) \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
+ apply (cases "abs i = 0", simp_all add: gcd_int_def)
apply (cases "abs j = 0", simp_all)
apply (cases "abs i = 1", simp_all)
apply (cases "abs j = 1", simp_all)
apply auto
done
lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow> m =0"
- by (induct t rule: numgcdh.induct, auto simp add:zgcd0)
+ by (induct t rule: numgcdh.induct, auto)
lemma dvdnumcoeff_aux:
assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
@@ -568,22 +568,22 @@
proof(induct t rule: numgcdh.induct)
case (2 n c t)
let ?g = "numgcdh t m"
- from prems have th:"zgcd c ?g > 1" by simp
+ from prems have th:"gcd c ?g > 1" by simp
from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
moreover {assume "abs c > 1" and gp: "?g > 1" with prems
have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)}
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp }
moreover {assume "abs c = 0 \<and> ?g > 1"
with prems have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp
hence ?case by simp }
moreover {assume "abs c > 1" and g0:"?g = 0"
from numgcdh0[OF g0] have "m=0". with prems have ?case by simp }
ultimately show ?case by blast
-qed(auto simp add: zgcd_zdvd1)
+qed auto
lemma dvdnumcoeff_aux2:
assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
@@ -727,7 +727,7 @@
constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
"simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
(let t' = simpnum t ; g = numgcd t' in
- if g > 1 then (let g' = zgcd n g in
+ if g > 1 then (let g' = gcd n g in
if g' = 1 then (t',n)
else (reducecoeffh t' g', n div g'))
else (t',n))))"
@@ -738,23 +738,23 @@
proof-
let ?t' = "simpnum t"
let ?g = "numgcd ?t'"
- let ?g' = "zgcd n ?g"
+ let ?g' = "gcd n ?g"
{assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
moreover
{ assume nnz: "n \<noteq> 0"
{assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)}
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
- from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith
+ from g1 nnz have gp0: "?g' \<noteq> 0" by simp
+ hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)}
moreover {assume g'1:"?g'>1"
from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
let ?tt = "reducecoeffh ?t' ?g'"
let ?t = "Inum bs ?tt"
- have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2)
- have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1)
+ have gpdg: "?g' dvd ?g" by simp
+ have gpdd: "?g' dvd n" by simp
have gpdgp: "?g' dvd ?g'" by simp
from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
have th2:"real ?g' * ?t = Inum bs ?t'" by simp
@@ -774,21 +774,21 @@
proof-
let ?t' = "simpnum t"
let ?g = "numgcd ?t'"
- let ?g' = "zgcd n ?g"
+ let ?g' = "gcd n ?g"
{assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)}
moreover
{ assume nnz: "n \<noteq> 0"
{assume "\<not> ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)}
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
- from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith
+ from g1 nnz have gp0: "?g' \<noteq> 0" by simp
+ hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis using prems
by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)}
moreover {assume g'1:"?g'>1"
- have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2)
- have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1)
+ have gpdg: "?g' dvd ?g" by simp
+ have gpdd: "?g' dvd n" by simp
have gpdgp: "?g' dvd ?g'" by simp
from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
--- a/src/HOL/Decision_Procs/MIR.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Tue Jun 23 15:48:55 2009 +0100
@@ -642,9 +642,9 @@
done
recdef numgcdh "measure size"
- "numgcdh (C i) = (\<lambda>g. zgcd i g)"
- "numgcdh (CN n c t) = (\<lambda>g. zgcd c (numgcdh t g))"
- "numgcdh (CF c s t) = (\<lambda>g. zgcd c (numgcdh t g))"
+ "numgcdh (C i) = (\<lambda>g. gcd i g)"
+ "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
+ "numgcdh (CF c s t) = (\<lambda>g. gcd c (numgcdh t g))"
"numgcdh t = (\<lambda>g. 1)"
definition
@@ -687,13 +687,13 @@
shows "Inum bs t = 0"
proof-
have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0"
- by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos)
+ by (induct t rule: numgcdh.induct, auto)
thus ?thesis using g0[simplified numgcd_def] by blast
qed
lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
using gp
- by (induct t rule: numgcdh.induct, auto simp add: zgcd_def)
+ by (induct t rule: numgcdh.induct, auto)
lemma numgcd_pos: "numgcd t \<ge>0"
by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
@@ -738,8 +738,8 @@
from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1)
qed simp_all
-lemma zgcd_gt1: "zgcd i j > 1 \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
- apply (unfold zgcd_def)
+lemma zgcd_gt1: "gcd i j > (1::int) \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
+ apply (unfold gcd_int_def)
apply (cases "i = 0", simp_all)
apply (cases "j = 0", simp_all)
apply (cases "abs i = 1", simp_all)
@@ -747,7 +747,7 @@
apply auto
done
lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow> m =0"
- by (induct t rule: numgcdh.induct, auto simp add:zgcd0)
+ by (induct t rule: numgcdh.induct, auto)
lemma dvdnumcoeff_aux:
assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
@@ -756,17 +756,17 @@
proof(induct t rule: numgcdh.induct)
case (2 n c t)
let ?g = "numgcdh t m"
- from prems have th:"zgcd c ?g > 1" by simp
+ from prems have th:"gcd c ?g > 1" by simp
from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
moreover {assume "abs c > 1" and gp: "?g > 1" with prems
have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)}
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp }
moreover {assume "abs c = 0 \<and> ?g > 1"
with prems have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp
hence ?case by simp }
moreover {assume "abs c > 1" and g0:"?g = 0"
from numgcdh0[OF g0] have "m=0". with prems have ?case by simp }
@@ -774,22 +774,22 @@
next
case (3 c s t)
let ?g = "numgcdh t m"
- from prems have th:"zgcd c ?g > 1" by simp
+ from prems have th:"gcd c ?g > 1" by simp
from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
moreover {assume "abs c > 1" and gp: "?g > 1" with prems
have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)}
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp }
moreover {assume "abs c = 0 \<and> ?g > 1"
with prems have th: "dvdnumcoeff t ?g" by simp
- have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2)
- from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)
+ have th': "gcd c ?g dvd ?g" by simp
+ from dvdnumcoeff_trans[OF th' th] have ?case by simp
hence ?case by simp }
moreover {assume "abs c > 1" and g0:"?g = 0"
from numgcdh0[OF g0] have "m=0". with prems have ?case by simp }
ultimately show ?case by blast
-qed(auto simp add: zgcd_zdvd1)
+qed auto
lemma dvdnumcoeff_aux2:
assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
@@ -1041,7 +1041,7 @@
constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
"simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
(let t' = simpnum t ; g = numgcd t' in
- if g > 1 then (let g' = zgcd n g in
+ if g > 1 then (let g' = gcd n g in
if g' = 1 then (t',n)
else (reducecoeffh t' g', n div g'))
else (t',n))))"
@@ -1052,23 +1052,23 @@
proof-
let ?t' = "simpnum t"
let ?g = "numgcd ?t'"
- let ?g' = "zgcd n ?g"
+ let ?g' = "gcd n ?g"
{assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
moreover
{ assume nnz: "n \<noteq> 0"
{assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
- from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith
+ from g1 nnz have gp0: "?g' \<noteq> 0" by simp
+ hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
moreover {assume g'1:"?g'>1"
from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
let ?tt = "reducecoeffh ?t' ?g'"
let ?t = "Inum bs ?tt"
- have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2)
- have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1)
+ have gpdg: "?g' dvd ?g" by simp
+ have gpdd: "?g' dvd n" by simp
have gpdgp: "?g' dvd ?g'" by simp
from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
have th2:"real ?g' * ?t = Inum bs ?t'" by simp
@@ -1088,21 +1088,21 @@
proof-
let ?t' = "simpnum t"
let ?g = "numgcd ?t'"
- let ?g' = "zgcd n ?g"
+ let ?g' = "gcd n ?g"
{assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)}
moreover
{ assume nnz: "n \<noteq> 0"
{assume "\<not> ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def)}
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
- from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith
+ from g1 nnz have gp0: "?g' \<noteq> 0" by simp
+ hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis using prems
by (auto simp add: Let_def simp_num_pair_def)}
moreover {assume g'1:"?g'>1"
- have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2)
- have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1)
+ have gpdg: "?g' dvd ?g" by simp
+ have gpdd: "?g' dvd n" by simp
have gpdgp: "?g' dvd ?g'" by simp
from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
@@ -1219,7 +1219,7 @@
constdefs simpdvd:: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)"
"simpdvd d t \<equiv>
(let g = numgcd t in
- if g > 1 then (let g' = zgcd d g in
+ if g > 1 then (let g' = gcd d g in
if g' = 1 then (d, t)
else (d div g',reducecoeffh t g'))
else (d, t))"
@@ -1228,20 +1228,20 @@
shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)"
proof-
let ?g = "numgcd t"
- let ?g' = "zgcd d ?g"
+ let ?g' = "gcd d ?g"
{assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)}
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
- from zgcd0 g1 dnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using zgcd_pos[where i="d" and j="numgcd t"] by arith
+ from g1 dnz have gp0: "?g' \<noteq> 0" by simp
+ hence g'p: "?g' > 0" using int_gcd_ge_0[where x="d" and y="numgcd t"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
moreover {assume g'1:"?g'>1"
from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" ..
let ?tt = "reducecoeffh t ?g'"
let ?t = "Inum bs ?tt"
- have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2)
- have gpdd: "?g' dvd d" by (simp add: zgcd_zdvd1)
+ have gpdg: "?g' dvd ?g" by simp
+ have gpdd: "?g' dvd d" by simp
have gpdgp: "?g' dvd ?g'" by simp
from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
have th2:"real ?g' * ?t = Inum bs t" by simp
@@ -2093,8 +2093,8 @@
"plusinf p = p"
recdef \<delta> "measure size"
- "\<delta> (And p q) = zlcm (\<delta> p) (\<delta> q)"
- "\<delta> (Or p q) = zlcm (\<delta> p) (\<delta> q)"
+ "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
+ "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
"\<delta> (Dvd i (CN 0 c e)) = i"
"\<delta> (NDvd i (CN 0 c e)) = i"
"\<delta> p = 1"
@@ -2126,21 +2126,21 @@
proof (induct p rule: iszlfm.induct)
case (1 p q)
let ?d = "\<delta> (And p q)"
- from prems zlcm_pos have dp: "?d >0" by simp
+ from prems int_lcm_pos have dp: "?d >0" by simp
have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp
hence th: "d\<delta> p ?d"
- using delta_mono prems by (auto simp del: dvd_zlcm_self1)
+ using delta_mono prems by(simp only: iszlfm.simps) blast
have "\<delta> q dvd \<delta> (And p q)" using prems by simp
- hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2)
+ hence th': "d\<delta> q ?d" using delta_mono prems by(simp only: iszlfm.simps) blast
from th th' dp show ?case by simp
next
case (2 p q)
let ?d = "\<delta> (And p q)"
- from prems zlcm_pos have dp: "?d >0" by simp
+ from prems int_lcm_pos have dp: "?d >0" by simp
have "\<delta> p dvd \<delta> (And p q)" using prems by simp hence th: "d\<delta> p ?d" using delta_mono prems
- by (auto simp del: dvd_zlcm_self1)
- have "\<delta> q dvd \<delta> (And p q)" using prems by simp hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2)
- from th th' dp show ?case by simp
+ by(simp only: iszlfm.simps) blast
+ have "\<delta> q dvd \<delta> (And p q)" using prems by simp hence th': "d\<delta> q ?d" using delta_mono prems by(simp only: iszlfm.simps) blast
+ from th th' dp show ?case by simp
qed simp_all
@@ -2388,8 +2388,8 @@
"d\<beta> p = (\<lambda> k. True)"
recdef \<zeta> "measure size"
- "\<zeta> (And p q) = zlcm (\<zeta> p) (\<zeta> q)"
- "\<zeta> (Or p q) = zlcm (\<zeta> p) (\<zeta> q)"
+ "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
+ "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
"\<zeta> (Eq (CN 0 c e)) = c"
"\<zeta> (NEq (CN 0 c e)) = c"
"\<zeta> (Lt (CN 0 c e)) = c"
@@ -2510,19 +2510,19 @@
using linp
proof(induct p rule: iszlfm.induct)
case (1 p q)
- from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: zlcm_pos)
+ from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
next
case (2 p q)
- from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp
- from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: zlcm_pos)
-qed (auto simp add: zlcm_pos)
+ from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
+ from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
+ dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
+qed (auto simp add: int_lcm_pos)
lemma a\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d\<beta> p l" and lp: "l > 0"
shows "iszlfm (a\<beta> p l) (a #bs) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm (real (l * x) #bs) (a\<beta> p l) = Ifm ((real x)#bs) p)"
@@ -4173,9 +4173,9 @@
lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
by (induct p rule: isrlfm.induct, auto)
-lemma zgcd_le1: assumes ip: "0 < i" shows "zgcd i j \<le> i"
+lemma zgcd_le1: assumes ip: "(0::int) < i" shows "gcd i j \<le> i"
proof-
- from zgcd_zdvd1 have th: "zgcd i j dvd i" by blast
+ from int_gcd_dvd1 have th: "gcd i j dvd i" by blast
from zdvd_imp_le[OF th ip] show ?thesis .
qed
@@ -5119,9 +5119,9 @@
let ?M = "?I x (minusinf ?rq)"
let ?P = "?I x (plusinf ?rq)"
have MF: "?M = False"
- apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def)
+ apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
- have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def)
+ have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
have "(\<exists> x. ?I x ?q ) =
((?I x (minusinf ?rq)) \<or> (?I x (plusinf ?rq )) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> ?rq). \<exists> (s,m) \<in> set (\<Upsilon> ?rq ). ?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))))"
@@ -5286,7 +5286,7 @@
let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))"
let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
from rlfm_l[OF qf] have lq: "isrlfm ?q"
- by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def zgcd_def)
+ by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def)
from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp
from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
from U_l UpU
--- a/src/HOL/Finite_Set.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Finite_Set.thy Tue Jun 23 15:48:55 2009 +0100
@@ -285,6 +285,10 @@
-- {* The image of a finite set is finite. *}
by (induct set: finite) simp_all
+lemma finite_image_set [simp]:
+ "finite {x. P x} \<Longrightarrow> finite { f x | x. P x }"
+ by (simp add: image_Collect [symmetric])
+
lemma finite_surj: "finite A ==> B <= f ` A ==> finite B"
apply (frule finite_imageI)
apply (erule finite_subset, assumption)
--- a/src/HOL/Fun.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Fun.thy Tue Jun 23 15:48:55 2009 +0100
@@ -133,7 +133,7 @@
shows "inj f"
using assms unfolding inj_on_def by auto
-text{*For Proofs in @{text "Tools/datatype_package/datatype_rep_proofs"}*}
+text{*For Proofs in @{text "Tools/Datatype/datatype_rep_proofs"}*}
lemma datatype_injI:
"(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)"
by (simp add: inj_on_def)
--- a/src/HOL/FunDef.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/FunDef.thy Tue Jun 23 15:48:55 2009 +0100
@@ -9,25 +9,25 @@
uses
"Tools/prop_logic.ML"
"Tools/sat_solver.ML"
- ("Tools/function_package/fundef_lib.ML")
- ("Tools/function_package/fundef_common.ML")
- ("Tools/function_package/inductive_wrap.ML")
- ("Tools/function_package/context_tree.ML")
- ("Tools/function_package/fundef_core.ML")
- ("Tools/function_package/sum_tree.ML")
- ("Tools/function_package/mutual.ML")
- ("Tools/function_package/pattern_split.ML")
- ("Tools/function_package/fundef_package.ML")
- ("Tools/function_package/auto_term.ML")
- ("Tools/function_package/measure_functions.ML")
- ("Tools/function_package/lexicographic_order.ML")
- ("Tools/function_package/fundef_datatype.ML")
- ("Tools/function_package/induction_scheme.ML")
- ("Tools/function_package/termination.ML")
- ("Tools/function_package/decompose.ML")
- ("Tools/function_package/descent.ML")
- ("Tools/function_package/scnp_solve.ML")
- ("Tools/function_package/scnp_reconstruct.ML")
+ ("Tools/Function/fundef_lib.ML")
+ ("Tools/Function/fundef_common.ML")
+ ("Tools/Function/inductive_wrap.ML")
+ ("Tools/Function/context_tree.ML")
+ ("Tools/Function/fundef_core.ML")
+ ("Tools/Function/sum_tree.ML")
+ ("Tools/Function/mutual.ML")
+ ("Tools/Function/pattern_split.ML")
+ ("Tools/Function/fundef.ML")
+ ("Tools/Function/auto_term.ML")
+ ("Tools/Function/measure_functions.ML")
+ ("Tools/Function/lexicographic_order.ML")
+ ("Tools/Function/fundef_datatype.ML")
+ ("Tools/Function/induction_scheme.ML")
+ ("Tools/Function/termination.ML")
+ ("Tools/Function/decompose.ML")
+ ("Tools/Function/descent.ML")
+ ("Tools/Function/scnp_solve.ML")
+ ("Tools/Function/scnp_reconstruct.ML")
begin
subsection {* Definitions with default value. *}
@@ -103,21 +103,21 @@
"wf R \<Longrightarrow> wfP (in_rel R)"
by (simp add: wfP_def)
-use "Tools/function_package/fundef_lib.ML"
-use "Tools/function_package/fundef_common.ML"
-use "Tools/function_package/inductive_wrap.ML"
-use "Tools/function_package/context_tree.ML"
-use "Tools/function_package/fundef_core.ML"
-use "Tools/function_package/sum_tree.ML"
-use "Tools/function_package/mutual.ML"
-use "Tools/function_package/pattern_split.ML"
-use "Tools/function_package/auto_term.ML"
-use "Tools/function_package/fundef_package.ML"
-use "Tools/function_package/fundef_datatype.ML"
-use "Tools/function_package/induction_scheme.ML"
+use "Tools/Function/fundef_lib.ML"
+use "Tools/Function/fundef_common.ML"
+use "Tools/Function/inductive_wrap.ML"
+use "Tools/Function/context_tree.ML"
+use "Tools/Function/fundef_core.ML"
+use "Tools/Function/sum_tree.ML"
+use "Tools/Function/mutual.ML"
+use "Tools/Function/pattern_split.ML"
+use "Tools/Function/auto_term.ML"
+use "Tools/Function/fundef.ML"
+use "Tools/Function/fundef_datatype.ML"
+use "Tools/Function/induction_scheme.ML"
setup {*
- FundefPackage.setup
+ Fundef.setup
#> FundefDatatype.setup
#> InductionScheme.setup
*}
@@ -127,7 +127,7 @@
inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
where is_measure_trivial: "is_measure f"
-use "Tools/function_package/measure_functions.ML"
+use "Tools/Function/measure_functions.ML"
setup MeasureFunctions.setup
lemma measure_size[measure_function]: "is_measure size"
@@ -138,7 +138,7 @@
lemma measure_snd[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (snd p))"
by (rule is_measure_trivial)
-use "Tools/function_package/lexicographic_order.ML"
+use "Tools/Function/lexicographic_order.ML"
setup LexicographicOrder.setup
@@ -307,11 +307,11 @@
subsection {* Tool setup *}
-use "Tools/function_package/termination.ML"
-use "Tools/function_package/decompose.ML"
-use "Tools/function_package/descent.ML"
-use "Tools/function_package/scnp_solve.ML"
-use "Tools/function_package/scnp_reconstruct.ML"
+use "Tools/Function/termination.ML"
+use "Tools/Function/decompose.ML"
+use "Tools/Function/descent.ML"
+use "Tools/Function/scnp_solve.ML"
+use "Tools/Function/scnp_reconstruct.ML"
setup {* ScnpReconstruct.setup *}
--- a/src/HOL/GCD.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/GCD.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,202 +1,604 @@
-(* Title: HOL/GCD.thy
- Author: Christophe Tabacznyj and Lawrence C Paulson
- Copyright 1996 University of Cambridge
+(* Title: GCD.thy
+ Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+ Thomas M. Rasmussen, Jeremy Avigad
+
+
+This file deals with the functions gcd and lcm, and properties of
+primes. Definitions and lemmas are proved uniformly for the natural
+numbers and integers.
+
+This file combines and revises a number of prior developments.
+
+The original theories "GCD" and "Primes" were by Christophe Tabacznyj
+and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
+gcd, lcm, and prime for the natural numbers.
+
+The original theory "IntPrimes" was by Thomas M. Rasmussen, and
+extended gcd, lcm, primes to the integers. Amine Chaieb provided
+another extension of the notions to the integers, and added a number
+of results to "Primes" and "GCD". IntPrimes also defined and developed
+the congruence relations on the integers. The notion was extended to
+the natural numbers by Chiaeb.
+
*)
-header {* The Greatest Common Divisor *}
+
+header {* GCD *}
theory GCD
-imports Main
+imports NatTransfer
+begin
+
+declare One_nat_def [simp del]
+
+subsection {* gcd *}
+
+class gcd = one +
+
+fixes
+ gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" and
+ lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+
begin
-text {*
- See \cite{davenport92}. \bigskip
-*}
+abbreviation
+ coprime :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "coprime x y == (gcd x y = 1)"
+
+end
+
+class prime = one +
+
+fixes
+ prime :: "'a \<Rightarrow> bool"
+
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: gcd
+
+begin
-subsection {* Specification of GCD on nats *}
+fun
+ gcd_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+ "gcd_nat x y =
+ (if y = 0 then x else gcd y (x mod y))"
+
+definition
+ lcm_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+ "lcm_nat x y = x * y div (gcd x y)"
+
+instance proof qed
+
+end
+
+
+instantiation nat :: prime
+
+begin
definition
- is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
- [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
- (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
+ prime_nat :: "nat \<Rightarrow> bool"
+where
+ [code del]: "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
+
+instance proof qed
-text {* Uniqueness *}
+end
+
-lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
- by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
+(* definitions for the integers *)
+
+instantiation int :: gcd
-text {* Connection to divides relation *}
+begin
-lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
- by (auto simp add: is_gcd_def)
+definition
+ gcd_int :: "int \<Rightarrow> int \<Rightarrow> int"
+where
+ "gcd_int x y = int (gcd (nat (abs x)) (nat (abs y)))"
-text {* Commutativity *}
+definition
+ lcm_int :: "int \<Rightarrow> int \<Rightarrow> int"
+where
+ "lcm_int x y = int (lcm (nat (abs x)) (nat (abs y)))"
-lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
- by (auto simp add: is_gcd_def)
+instance proof qed
+
+end
-subsection {* GCD on nat by Euclid's algorithm *}
+instantiation int :: prime
+
+begin
+
+definition
+ prime_int :: "int \<Rightarrow> bool"
+where
+ [code del]: "prime_int p = prime (nat p)"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_gcd:
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> gcd (nat x) (nat y) = nat (gcd x y)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> lcm (nat x) (nat y) = nat (lcm x y)"
+ "(x::int) >= 0 \<Longrightarrow> prime (nat x) = prime x"
+ unfolding gcd_int_def lcm_int_def prime_int_def
+ by auto
-fun
- gcd :: "nat => nat => nat"
-where
- "gcd m n = (if n = 0 then m else gcd n (m mod n))"
-lemma gcd_induct [case_names "0" rec]:
+lemma transfer_nat_int_gcd_closures:
+ "x >= (0::int) \<Longrightarrow> y >= 0 \<Longrightarrow> gcd x y >= 0"
+ "x >= (0::int) \<Longrightarrow> y >= 0 \<Longrightarrow> lcm x y >= 0"
+ by (auto simp add: gcd_int_def lcm_int_def)
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_gcd transfer_nat_int_gcd_closures]
+
+lemma transfer_int_nat_gcd:
+ "gcd (int x) (int y) = int (gcd x y)"
+ "lcm (int x) (int y) = int (lcm x y)"
+ "prime (int x) = prime x"
+ by (unfold gcd_int_def lcm_int_def prime_int_def, auto)
+
+lemma transfer_int_nat_gcd_closures:
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> gcd x y >= 0"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> lcm x y >= 0"
+ by (auto simp add: gcd_int_def lcm_int_def)
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_gcd transfer_int_nat_gcd_closures]
+
+
+
+subsection {* GCD *}
+
+(* was gcd_induct *)
+lemma nat_gcd_induct:
fixes m n :: nat
assumes "\<And>m. P m 0"
and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
shows "P m n"
-proof (induct m n rule: gcd.induct)
- case (1 m n) with assms show ?case by (cases "n = 0") simp_all
-qed
+ apply (rule gcd_nat.induct)
+ apply (case_tac "y = 0")
+ using assms apply simp_all
+done
+
+(* specific to int *)
+
+lemma int_gcd_neg1 [simp]: "gcd (-x::int) y = gcd x y"
+ by (simp add: gcd_int_def)
+
+lemma int_gcd_neg2 [simp]: "gcd (x::int) (-y) = gcd x y"
+ by (simp add: gcd_int_def)
+
+lemma int_gcd_abs: "gcd (x::int) y = gcd (abs x) (abs y)"
+ by (simp add: gcd_int_def)
+
+lemma int_gcd_cases:
+ fixes x :: int and y
+ assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd x y)"
+ and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd x (-y))"
+ and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd (-x) y)"
+ and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd (-x) (-y))"
+ shows "P (gcd x y)"
+by (insert prems, auto simp add: int_gcd_neg1 int_gcd_neg2, arith)
-lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
- by simp
+lemma int_gcd_ge_0 [simp]: "gcd (x::int) y >= 0"
+ by (simp add: gcd_int_def)
+
+lemma int_lcm_neg1: "lcm (-x::int) y = lcm x y"
+ by (simp add: lcm_int_def)
+
+lemma int_lcm_neg2: "lcm (x::int) (-y) = lcm x y"
+ by (simp add: lcm_int_def)
+
+lemma int_lcm_abs: "lcm (x::int) y = lcm (abs x) (abs y)"
+ by (simp add: lcm_int_def)
-lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
+lemma int_lcm_cases:
+ fixes x :: int and y
+ assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm x y)"
+ and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm x (-y))"
+ and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm (-x) y)"
+ and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm (-x) (-y))"
+ shows "P (lcm x y)"
+by (insert prems, auto simp add: int_lcm_neg1 int_lcm_neg2, arith)
+
+lemma int_lcm_ge_0 [simp]: "lcm (x::int) y >= 0"
+ by (simp add: lcm_int_def)
+
+(* was gcd_0, etc. *)
+lemma nat_gcd_0 [simp]: "gcd (x::nat) 0 = x"
by simp
-lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
+(* was igcd_0, etc. *)
+lemma int_gcd_0 [simp]: "gcd (x::int) 0 = abs x"
+ by (unfold gcd_int_def, auto)
+
+lemma nat_gcd_0_left [simp]: "gcd 0 (x::nat) = x"
by simp
-lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
+lemma int_gcd_0_left [simp]: "gcd 0 (x::int) = abs x"
+ by (unfold gcd_int_def, auto)
+
+lemma nat_gcd_red: "gcd (x::nat) y = gcd y (x mod y)"
+ by (case_tac "y = 0", auto)
+
+(* weaker, but useful for the simplifier *)
+
+lemma nat_gcd_non_0: "y ~= (0::nat) \<Longrightarrow> gcd (x::nat) y = gcd y (x mod y)"
+ by simp
+
+lemma nat_gcd_1 [simp]: "gcd (m::nat) 1 = 1"
by simp
-lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
- unfolding One_nat_def by (rule gcd_1)
+lemma nat_gcd_Suc_0 [simp]: "gcd (m::nat) (Suc 0) = Suc 0"
+ by (simp add: One_nat_def)
+
+lemma int_gcd_1 [simp]: "gcd (m::int) 1 = 1"
+ by (simp add: gcd_int_def)
-declare gcd.simps [simp del]
+lemma nat_gcd_self [simp]: "gcd (x::nat) x = x"
+ by simp
+
+lemma int_gcd_def [simp]: "gcd (x::int) x = abs x"
+ by (subst int_gcd_abs, auto simp add: gcd_int_def)
+
+declare gcd_nat.simps [simp del]
text {*
\medskip @{term "gcd m n"} divides @{text m} and @{text n}. The
conjunctions don't seem provable separately.
*}
-lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
- and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
- apply (induct m n rule: gcd_induct)
- apply (simp_all add: gcd_non_0)
+lemma nat_gcd_dvd1 [iff]: "(gcd (m::nat)) n dvd m"
+ and nat_gcd_dvd2 [iff]: "(gcd m n) dvd n"
+ apply (induct m n rule: nat_gcd_induct)
+ apply (simp_all add: nat_gcd_non_0)
apply (blast dest: dvd_mod_imp_dvd)
- done
+done
+
+thm nat_gcd_dvd1 [transferred]
+
+lemma int_gcd_dvd1 [iff]: "gcd (x::int) y dvd x"
+ apply (subst int_gcd_abs)
+ apply (rule dvd_trans)
+ apply (rule nat_gcd_dvd1 [transferred])
+ apply auto
+done
-text {*
- \medskip Maximality: for all @{term m}, @{term n}, @{term k}
- naturals, if @{term k} divides @{term m} and @{term k} divides
- @{term n} then @{term k} divides @{term "gcd m n"}.
-*}
+lemma int_gcd_dvd2 [iff]: "gcd (x::int) y dvd y"
+ apply (subst int_gcd_abs)
+ apply (rule dvd_trans)
+ apply (rule nat_gcd_dvd2 [transferred])
+ apply auto
+done
+
+lemma dvd_gcd_D1_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd m"
+by(metis nat_gcd_dvd1 dvd_trans)
+
+lemma dvd_gcd_D2_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd n"
+by(metis nat_gcd_dvd2 dvd_trans)
+
+lemma dvd_gcd_D1_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd m"
+by(metis int_gcd_dvd1 dvd_trans)
-lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
- by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
+lemma dvd_gcd_D2_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd n"
+by(metis int_gcd_dvd2 dvd_trans)
+
+lemma nat_gcd_le1 [simp]: "a \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> a"
+ by (rule dvd_imp_le, auto)
+
+lemma nat_gcd_le2 [simp]: "b \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> b"
+ by (rule dvd_imp_le, auto)
+
+lemma int_gcd_le1 [simp]: "a > 0 \<Longrightarrow> gcd (a::int) b \<le> a"
+ by (rule zdvd_imp_le, auto)
+
+lemma int_gcd_le2 [simp]: "b > 0 \<Longrightarrow> gcd (a::int) b \<le> b"
+ by (rule zdvd_imp_le, auto)
-text {*
- \medskip Function gcd yields the Greatest Common Divisor.
-*}
+lemma nat_gcd_greatest: "(k::nat) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
+ by (induct m n rule: nat_gcd_induct) (simp_all add: nat_gcd_non_0 dvd_mod)
+
+lemma int_gcd_greatest:
+ assumes "(k::int) dvd m" and "k dvd n"
+ shows "k dvd gcd m n"
-lemma is_gcd: "is_gcd m n (gcd m n) "
- by (simp add: is_gcd_def gcd_greatest)
+ apply (subst int_gcd_abs)
+ apply (subst abs_dvd_iff [symmetric])
+ apply (rule nat_gcd_greatest [transferred])
+ using prems apply auto
+done
+lemma nat_gcd_greatest_iff [iff]: "(k dvd gcd (m::nat) n) =
+ (k dvd m & k dvd n)"
+ by (blast intro!: nat_gcd_greatest intro: dvd_trans)
-subsection {* Derived laws for GCD *}
+lemma int_gcd_greatest_iff: "((k::int) dvd gcd m n) = (k dvd m & k dvd n)"
+ by (blast intro!: int_gcd_greatest intro: dvd_trans)
-lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
- by (blast intro!: gcd_greatest intro: dvd_trans)
+lemma nat_gcd_zero [simp]: "(gcd (m::nat) n = 0) = (m = 0 & n = 0)"
+ by (simp only: dvd_0_left_iff [symmetric] nat_gcd_greatest_iff)
+
+lemma int_gcd_zero [simp]: "(gcd (m::int) n = 0) = (m = 0 & n = 0)"
+ by (auto simp add: gcd_int_def)
-lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
- by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
+lemma nat_gcd_pos [simp]: "(gcd (m::nat) n > 0) = (m ~= 0 | n ~= 0)"
+ by (insert nat_gcd_zero [of m n], arith)
+
+lemma int_gcd_pos [simp]: "(gcd (m::int) n > 0) = (m ~= 0 | n ~= 0)"
+ by (insert int_gcd_zero [of m n], insert int_gcd_ge_0 [of m n], arith)
+
+lemma nat_gcd_commute: "gcd (m::nat) n = gcd n m"
+ by (rule dvd_anti_sym, auto)
-lemma gcd_commute: "gcd m n = gcd n m"
- apply (rule is_gcd_unique)
- apply (rule is_gcd)
- apply (subst is_gcd_commute)
- apply (simp add: is_gcd)
- done
+lemma int_gcd_commute: "gcd (m::int) n = gcd n m"
+ by (auto simp add: gcd_int_def nat_gcd_commute)
+
+lemma nat_gcd_assoc: "gcd (gcd (k::nat) m) n = gcd k (gcd m n)"
+ apply (rule dvd_anti_sym)
+ apply (blast intro: dvd_trans)+
+done
+
+lemma int_gcd_assoc: "gcd (gcd (k::int) m) n = gcd k (gcd m n)"
+ by (auto simp add: gcd_int_def nat_gcd_assoc)
+
+lemmas nat_gcd_left_commute =
+ mk_left_commute[of gcd, OF nat_gcd_assoc nat_gcd_commute]
+
+lemmas int_gcd_left_commute =
+ mk_left_commute[of gcd, OF int_gcd_assoc int_gcd_commute]
-lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
- apply (rule is_gcd_unique)
- apply (rule is_gcd)
- apply (simp add: is_gcd_def)
- apply (blast intro: dvd_trans)
- done
+lemmas nat_gcd_ac = nat_gcd_assoc nat_gcd_commute nat_gcd_left_commute
+ -- {* gcd is an AC-operator *}
+
+lemmas int_gcd_ac = int_gcd_assoc int_gcd_commute int_gcd_left_commute
+
+lemma nat_gcd_1_left [simp]: "gcd (1::nat) m = 1"
+ by (subst nat_gcd_commute, simp)
+
+lemma nat_gcd_Suc_0_left [simp]: "gcd (Suc 0) m = Suc 0"
+ by (subst nat_gcd_commute, simp add: One_nat_def)
+
+lemma int_gcd_1_left [simp]: "gcd (1::int) m = 1"
+ by (subst int_gcd_commute, simp)
-lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
- by (simp add: gcd_commute)
+lemma nat_gcd_unique: "(d::nat) dvd a \<and> d dvd b \<and>
+ (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
+ apply auto
+ apply (rule dvd_anti_sym)
+ apply (erule (1) nat_gcd_greatest)
+ apply auto
+done
-lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
- unfolding One_nat_def by (rule gcd_1_left)
+lemma int_gcd_unique: "d >= 0 & (d::int) dvd a \<and> d dvd b \<and>
+ (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
+ apply (case_tac "d = 0")
+ apply force
+ apply (rule iffI)
+ apply (rule zdvd_anti_sym)
+ apply arith
+ apply (subst int_gcd_pos)
+ apply clarsimp
+ apply (drule_tac x = "d + 1" in spec)
+ apply (frule zdvd_imp_le)
+ apply (auto intro: int_gcd_greatest)
+done
text {*
\medskip Multiplication laws
*}
-lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
+lemma nat_gcd_mult_distrib: "(k::nat) * gcd m n = gcd (k * m) (k * n)"
-- {* \cite[page 27]{davenport92} *}
- apply (induct m n rule: gcd_induct)
- apply simp
+ apply (induct m n rule: nat_gcd_induct)
+ apply simp
apply (case_tac "k = 0")
- apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
- done
+ apply (simp_all add: mod_geq nat_gcd_non_0 mod_mult_distrib2)
+done
-lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
- apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
- done
+lemma int_gcd_mult_distrib: "abs (k::int) * gcd m n = gcd (k * m) (k * n)"
+ apply (subst (1 2) int_gcd_abs)
+ apply (simp add: abs_mult)
+ apply (rule nat_gcd_mult_distrib [transferred])
+ apply auto
+done
-lemma gcd_self [simp, algebra]: "gcd k k = k"
- apply (rule gcd_mult [of k 1, simplified])
- done
+lemma nat_gcd_mult [simp]: "gcd (k::nat) (k * n) = k"
+ by (rule nat_gcd_mult_distrib [of k 1 n, simplified, symmetric])
-lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
- apply (insert gcd_mult_distrib2 [of m k n])
+lemma int_gcd_mult [simp]: "gcd (k::int) (k * n) = abs k"
+ by (rule int_gcd_mult_distrib [of k 1 n, simplified, symmetric])
+
+lemma nat_coprime_dvd_mult: "coprime (k::nat) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
+ apply (insert nat_gcd_mult_distrib [of m k n])
apply simp
apply (erule_tac t = m in ssubst)
apply simp
done
-lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
- by (auto intro: relprime_dvd_mult dvd_mult2)
+lemma int_coprime_dvd_mult:
+ assumes "coprime (k::int) n" and "k dvd m * n"
+ shows "k dvd m"
-lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
+ using prems
+ apply (subst abs_dvd_iff [symmetric])
+ apply (subst dvd_abs_iff [symmetric])
+ apply (subst (asm) int_gcd_abs)
+ apply (rule nat_coprime_dvd_mult [transferred])
+ apply auto
+ apply (subst abs_mult [symmetric], auto)
+done
+
+lemma nat_coprime_dvd_mult_iff: "coprime (k::nat) n \<Longrightarrow>
+ (k dvd m * n) = (k dvd m)"
+ by (auto intro: nat_coprime_dvd_mult)
+
+lemma int_coprime_dvd_mult_iff: "coprime (k::int) n \<Longrightarrow>
+ (k dvd m * n) = (k dvd m)"
+ by (auto intro: int_coprime_dvd_mult)
+
+lemma nat_gcd_mult_cancel: "coprime k n \<Longrightarrow> gcd ((k::nat) * m) n = gcd m n"
apply (rule dvd_anti_sym)
- apply (rule gcd_greatest)
- apply (rule_tac n = k in relprime_dvd_mult)
- apply (simp add: gcd_assoc)
- apply (simp add: gcd_commute)
- apply (simp_all add: mult_commute)
- done
+ apply (rule nat_gcd_greatest)
+ apply (rule_tac n = k in nat_coprime_dvd_mult)
+ apply (simp add: nat_gcd_assoc)
+ apply (simp add: nat_gcd_commute)
+ apply (simp_all add: mult_commute)
+done
+lemma int_gcd_mult_cancel:
+ assumes "coprime (k::int) n"
+ shows "gcd (k * m) n = gcd m n"
+
+ using prems
+ apply (subst (1 2) int_gcd_abs)
+ apply (subst abs_mult)
+ apply (rule nat_gcd_mult_cancel [transferred])
+ apply (auto simp add: int_gcd_abs [symmetric])
+done
text {* \medskip Addition laws *}
-lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
- by (cases "n = 0") (auto simp add: gcd_non_0)
+lemma nat_gcd_add1 [simp]: "gcd ((m::nat) + n) n = gcd m n"
+ apply (case_tac "n = 0")
+ apply (simp_all add: nat_gcd_non_0)
+done
+
+lemma nat_gcd_add2 [simp]: "gcd (m::nat) (m + n) = gcd m n"
+ apply (subst (1 2) nat_gcd_commute)
+ apply (subst add_commute)
+ apply simp
+done
+
+(* to do: add the other variations? *)
+
+lemma nat_gcd_diff1: "(m::nat) >= n \<Longrightarrow> gcd (m - n) n = gcd m n"
+ by (subst nat_gcd_add1 [symmetric], auto)
+
+lemma nat_gcd_diff2: "(n::nat) >= m \<Longrightarrow> gcd (n - m) n = gcd m n"
+ apply (subst nat_gcd_commute)
+ apply (subst nat_gcd_diff1 [symmetric])
+ apply auto
+ apply (subst nat_gcd_commute)
+ apply (subst nat_gcd_diff1)
+ apply assumption
+ apply (rule nat_gcd_commute)
+done
+
+lemma int_gcd_non_0: "(y::int) > 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
+ apply (frule_tac b = y and a = x in pos_mod_sign)
+ apply (simp del: pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib)
+ apply (auto simp add: nat_gcd_non_0 nat_mod_distrib [symmetric]
+ zmod_zminus1_eq_if)
+ apply (frule_tac a = x in pos_mod_bound)
+ apply (subst (1 2) nat_gcd_commute)
+ apply (simp del: pos_mod_bound add: nat_diff_distrib nat_gcd_diff2
+ nat_le_eq_zle)
+done
+
+lemma int_gcd_red: "gcd (x::int) y = gcd y (x mod y)"
+ apply (case_tac "y = 0")
+ apply force
+ apply (case_tac "y > 0")
+ apply (subst int_gcd_non_0, auto)
+ apply (insert int_gcd_non_0 [of "-y" "-x"])
+ apply (auto simp add: int_gcd_neg1 int_gcd_neg2)
+done
-lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
-proof -
- have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
- also have "... = gcd (n + m) m" by (simp add: add_commute)
- also have "... = gcd n m" by simp
- also have "... = gcd m n" by (rule gcd_commute)
- finally show ?thesis .
+lemma int_gcd_add1 [simp]: "gcd ((m::int) + n) n = gcd m n"
+ apply (case_tac "n = 0", force)
+ apply (subst (1 2) int_gcd_red)
+ apply auto
+done
+
+lemma int_gcd_add2 [simp]: "gcd m ((m::int) + n) = gcd m n"
+ apply (subst int_gcd_commute)
+ apply (subst add_commute)
+ apply (subst int_gcd_add1)
+ apply (subst int_gcd_commute)
+ apply (rule refl)
+done
+
+lemma nat_gcd_add_mult: "gcd (m::nat) (k * m + n) = gcd m n"
+ by (induct k, simp_all add: ring_simps)
+
+lemma int_gcd_add_mult: "gcd (m::int) (k * m + n) = gcd m n"
+ apply (subgoal_tac "ALL s. ALL m. ALL n.
+ gcd m (int (s::nat) * m + n) = gcd m n")
+ apply (case_tac "k >= 0")
+ apply (drule_tac x = "nat k" in spec, force)
+ apply (subst (1 2) int_gcd_neg2 [symmetric])
+ apply (drule_tac x = "nat (- k)" in spec)
+ apply (drule_tac x = "m" in spec)
+ apply (drule_tac x = "-n" in spec)
+ apply auto
+ apply (rule nat_induct)
+ apply auto
+ apply (auto simp add: left_distrib)
+ apply (subst add_assoc)
+ apply simp
+done
+
+(* to do: differences, and all variations of addition rules
+ as simplification rules for nat and int *)
+
+lemma nat_gcd_dvd_prod [iff]: "gcd (m::nat) n dvd k * n"
+ using mult_dvd_mono [of 1] by auto
+
+(* to do: add the three variations of these, and for ints? *)
+
+lemma finite_divisors_nat:
+ assumes "(m::nat)~= 0" shows "finite{d. d dvd m}"
+proof-
+ have "finite{d. d <= m}" by(blast intro: bounded_nat_set_is_finite)
+ from finite_subset[OF _ this] show ?thesis using assms
+ by(bestsimp intro!:dvd_imp_le)
qed
-lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
- apply (subst add_commute)
- apply (rule gcd_add2)
- done
-
-lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
- by (induct k) (simp_all add: add_assoc)
+lemma finite_divisors_int:
+ assumes "(i::int) ~= 0" shows "finite{d. d dvd i}"
+proof-
+ have "{d. abs d <= abs i} = {- abs i .. abs i}" by(auto simp:abs_if)
+ hence "finite{d. abs d <= abs i}" by simp
+ from finite_subset[OF _ this] show ?thesis using assms
+ by(bestsimp intro!:dvd_imp_le_int)
+qed
-lemma gcd_dvd_prod: "gcd m n dvd m * n"
- using mult_dvd_mono [of 1] by auto
+lemma gcd_is_Max_divisors_nat:
+ "m ~= 0 \<Longrightarrow> n ~= 0 \<Longrightarrow> gcd (m::nat) n = (Max {d. d dvd m & d dvd n})"
+apply(rule Max_eqI[THEN sym])
+ apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_nat)
+ apply simp
+ apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le nat_gcd_greatest_iff nat_gcd_pos)
+apply simp
+done
-text {*
- \medskip Division by gcd yields rrelatively primes.
-*}
+lemma gcd_is_Max_divisors_int:
+ "m ~= 0 ==> n ~= 0 ==> gcd (m::int) n = (Max {d. d dvd m & d dvd n})"
+apply(rule Max_eqI[THEN sym])
+ apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_int)
+ apply simp
+ apply (metis int_gcd_greatest_iff int_gcd_pos zdvd_imp_le)
+apply simp
+done
-lemma div_gcd_relprime:
- assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
- shows "gcd (a div gcd a b) (b div gcd a b) = 1"
+
+subsection {* Coprimality *}
+
+lemma nat_div_gcd_coprime [intro]:
+ assumes nz: "(a::nat) \<noteq> 0 \<or> b \<noteq> 0"
+ shows "coprime (a div gcd a b) (b div gcd a b)"
proof -
let ?g = "gcd a b"
let ?a' = "a div ?g"
@@ -207,42 +609,551 @@
from dvdg dvdg' obtain ka kb ka' kb' where
kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
unfolding dvd_def by blast
- then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
+ then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'"
+ by simp_all
then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
- have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
- then have gp: "?g > 0" by simp
- from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+ have "?g \<noteq> 0" using nz by (simp add: nat_gcd_zero)
+ then have gp: "?g > 0" by arith
+ from nat_gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
qed
+lemma int_div_gcd_coprime [intro]:
+ assumes nz: "(a::int) \<noteq> 0 \<or> b \<noteq> 0"
+ shows "coprime (a div gcd a b) (b div gcd a b)"
-lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
-proof(auto)
- assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
- from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b]
- have th: "gcd a b dvd d" by blast
- from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast
+ apply (subst (1 2 3) int_gcd_abs)
+ apply (subst (1 2) abs_div)
+ apply auto
+ prefer 3
+ apply (rule nat_div_gcd_coprime [transferred])
+ using nz apply (auto simp add: int_gcd_abs [symmetric])
+done
+
+lemma nat_coprime: "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+ using nat_gcd_unique[of 1 a b, simplified] by auto
+
+lemma nat_coprime_Suc_0:
+ "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = Suc 0)"
+ using nat_coprime by (simp add: One_nat_def)
+
+lemma int_coprime: "coprime (a::int) b \<longleftrightarrow>
+ (\<forall>d. d >= 0 \<and> d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+ using int_gcd_unique [of 1 a b]
+ apply clarsimp
+ apply (erule subst)
+ apply (rule iffI)
+ apply force
+ apply (drule_tac x = "abs e" in exI)
+ apply (case_tac "e >= 0")
+ apply force
+ apply force
+done
+
+lemma nat_gcd_coprime:
+ assumes z: "gcd (a::nat) b \<noteq> 0" and a: "a = a' * gcd a b" and
+ b: "b = b' * gcd a b"
+ shows "coprime a' b'"
+
+ apply (subgoal_tac "a' = a div gcd a b")
+ apply (erule ssubst)
+ apply (subgoal_tac "b' = b div gcd a b")
+ apply (erule ssubst)
+ apply (rule nat_div_gcd_coprime)
+ using prems
+ apply force
+ apply (subst (1) b)
+ using z apply force
+ apply (subst (1) a)
+ using z apply force
+done
+
+lemma int_gcd_coprime:
+ assumes z: "gcd (a::int) b \<noteq> 0" and a: "a = a' * gcd a b" and
+ b: "b = b' * gcd a b"
+ shows "coprime a' b'"
+
+ apply (subgoal_tac "a' = a div gcd a b")
+ apply (erule ssubst)
+ apply (subgoal_tac "b' = b div gcd a b")
+ apply (erule ssubst)
+ apply (rule int_div_gcd_coprime)
+ using prems
+ apply force
+ apply (subst (1) b)
+ using z apply force
+ apply (subst (1) a)
+ using z apply force
+done
+
+lemma nat_coprime_mult: assumes da: "coprime (d::nat) a" and db: "coprime d b"
+ shows "coprime d (a * b)"
+ apply (subst nat_gcd_commute)
+ using da apply (subst nat_gcd_mult_cancel)
+ apply (subst nat_gcd_commute, assumption)
+ apply (subst nat_gcd_commute, rule db)
+done
+
+lemma int_coprime_mult: assumes da: "coprime (d::int) a" and db: "coprime d b"
+ shows "coprime d (a * b)"
+ apply (subst int_gcd_commute)
+ using da apply (subst int_gcd_mult_cancel)
+ apply (subst int_gcd_commute, assumption)
+ apply (subst int_gcd_commute, rule db)
+done
+
+lemma nat_coprime_lmult:
+ assumes dab: "coprime (d::nat) (a * b)" shows "coprime d a"
+proof -
+ have "gcd d a dvd gcd d (a * b)"
+ by (rule nat_gcd_greatest, auto)
+ with dab show ?thesis
+ by auto
+qed
+
+lemma int_coprime_lmult:
+ assumes dab: "coprime (d::int) (a * b)" shows "coprime d a"
+proof -
+ have "gcd d a dvd gcd d (a * b)"
+ by (rule int_gcd_greatest, auto)
+ with dab show ?thesis
+ by auto
+qed
+
+lemma nat_coprime_rmult:
+ assumes dab: "coprime (d::nat) (a * b)" shows "coprime d b"
+proof -
+ have "gcd d b dvd gcd d (a * b)"
+ by (rule nat_gcd_greatest, auto intro: dvd_mult)
+ with dab show ?thesis
+ by auto
+qed
+
+lemma int_coprime_rmult:
+ assumes dab: "coprime (d::int) (a * b)" shows "coprime d b"
+proof -
+ have "gcd d b dvd gcd d (a * b)"
+ by (rule int_gcd_greatest, auto intro: dvd_mult)
+ with dab show ?thesis
+ by auto
+qed
+
+lemma nat_coprime_mul_eq: "coprime (d::nat) (a * b) \<longleftrightarrow>
+ coprime d a \<and> coprime d b"
+ using nat_coprime_rmult[of d a b] nat_coprime_lmult[of d a b]
+ nat_coprime_mult[of d a b]
+ by blast
+
+lemma int_coprime_mul_eq: "coprime (d::int) (a * b) \<longleftrightarrow>
+ coprime d a \<and> coprime d b"
+ using int_coprime_rmult[of d a b] int_coprime_lmult[of d a b]
+ int_coprime_mult[of d a b]
+ by blast
+
+lemma nat_gcd_coprime_exists:
+ assumes nz: "gcd (a::nat) b \<noteq> 0"
+ shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
+ apply (rule_tac x = "a div gcd a b" in exI)
+ apply (rule_tac x = "b div gcd a b" in exI)
+ using nz apply (auto simp add: nat_div_gcd_coprime dvd_div_mult)
+done
+
+lemma int_gcd_coprime_exists:
+ assumes nz: "gcd (a::int) b \<noteq> 0"
+ shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
+ apply (rule_tac x = "a div gcd a b" in exI)
+ apply (rule_tac x = "b div gcd a b" in exI)
+ using nz apply (auto simp add: int_div_gcd_coprime dvd_div_mult_self)
+done
+
+lemma nat_coprime_exp: "coprime (d::nat) a \<Longrightarrow> coprime d (a^n)"
+ by (induct n, simp_all add: nat_coprime_mult)
+
+lemma int_coprime_exp: "coprime (d::int) a \<Longrightarrow> coprime d (a^n)"
+ by (induct n, simp_all add: int_coprime_mult)
+
+lemma nat_coprime_exp2 [intro]: "coprime (a::nat) b \<Longrightarrow> coprime (a^n) (b^m)"
+ apply (rule nat_coprime_exp)
+ apply (subst nat_gcd_commute)
+ apply (rule nat_coprime_exp)
+ apply (subst nat_gcd_commute, assumption)
+done
+
+lemma int_coprime_exp2 [intro]: "coprime (a::int) b \<Longrightarrow> coprime (a^n) (b^m)"
+ apply (rule int_coprime_exp)
+ apply (subst int_gcd_commute)
+ apply (rule int_coprime_exp)
+ apply (subst int_gcd_commute, assumption)
+done
+
+lemma nat_gcd_exp: "gcd ((a::nat)^n) (b^n) = (gcd a b)^n"
+proof (cases)
+ assume "a = 0 & b = 0"
+ thus ?thesis by simp
+ next assume "~(a = 0 & b = 0)"
+ hence "coprime ((a div gcd a b)^n) ((b div gcd a b)^n)"
+ by auto
+ hence "gcd ((a div gcd a b)^n * (gcd a b)^n)
+ ((b div gcd a b)^n * (gcd a b)^n) = (gcd a b)^n"
+ apply (subst (1 2) mult_commute)
+ apply (subst nat_gcd_mult_distrib [symmetric])
+ apply simp
+ done
+ also have "(a div gcd a b)^n * (gcd a b)^n = a^n"
+ apply (subst div_power)
+ apply auto
+ apply (rule dvd_div_mult_self)
+ apply (rule dvd_power_same)
+ apply auto
+ done
+ also have "(b div gcd a b)^n * (gcd a b)^n = b^n"
+ apply (subst div_power)
+ apply auto
+ apply (rule dvd_div_mult_self)
+ apply (rule dvd_power_same)
+ apply auto
+ done
+ finally show ?thesis .
+qed
+
+lemma int_gcd_exp: "gcd ((a::int)^n) (b^n) = (gcd a b)^n"
+ apply (subst (1 2) int_gcd_abs)
+ apply (subst (1 2) power_abs)
+ apply (rule nat_gcd_exp [where n = n, transferred])
+ apply auto
+done
+
+lemma nat_coprime_divprod: "(d::nat) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
+ using nat_coprime_dvd_mult_iff[of d a b]
+ by (auto simp add: mult_commute)
+
+lemma int_coprime_divprod: "(d::int) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
+ using int_coprime_dvd_mult_iff[of d a b]
+ by (auto simp add: mult_commute)
+
+lemma nat_division_decomp: assumes dc: "(a::nat) dvd b * c"
+ shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
+proof-
+ let ?g = "gcd a b"
+ {assume "?g = 0" with dc have ?thesis by auto}
+ moreover
+ {assume z: "?g \<noteq> 0"
+ from nat_gcd_coprime_exists[OF z]
+ obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
+ by blast
+ have thb: "?g dvd b" by auto
+ from ab'(1) have "a' dvd a" unfolding dvd_def by blast
+ with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
+ from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
+ hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
+ with z have th_1: "a' dvd b' * c" by auto
+ from nat_coprime_dvd_mult[OF ab'(3)] th_1
+ have thc: "a' dvd c" by (subst (asm) mult_commute, blast)
+ from ab' have "a = ?g*a'" by algebra
+ with thb thc have ?thesis by blast }
+ ultimately show ?thesis by blast
+qed
+
+lemma int_division_decomp: assumes dc: "(a::int) dvd b * c"
+ shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
+proof-
+ let ?g = "gcd a b"
+ {assume "?g = 0" with dc have ?thesis by auto}
+ moreover
+ {assume z: "?g \<noteq> 0"
+ from int_gcd_coprime_exists[OF z]
+ obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
+ by blast
+ have thb: "?g dvd b" by auto
+ from ab'(1) have "a' dvd a" unfolding dvd_def by blast
+ with dc have th0: "a' dvd b*c"
+ using dvd_trans[of a' a "b*c"] by simp
+ from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
+ hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
+ with z have th_1: "a' dvd b' * c" by auto
+ from int_coprime_dvd_mult[OF ab'(3)] th_1
+ have thc: "a' dvd c" by (subst (asm) mult_commute, blast)
+ from ab' have "a = ?g*a'" by algebra
+ with thb thc have ?thesis by blast }
+ ultimately show ?thesis by blast
qed
-lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
- shows "gcd x y = gcd u v"
+lemma nat_pow_divides_pow:
+ assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0"
+ shows "a dvd b"
+proof-
+ let ?g = "gcd a b"
+ from n obtain m where m: "n = Suc m" by (cases n, simp_all)
+ {assume "?g = 0" with ab n have ?thesis by auto }
+ moreover
+ {assume z: "?g \<noteq> 0"
+ hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
+ from nat_gcd_coprime_exists[OF z]
+ obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
+ by blast
+ from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n"
+ by (simp add: ab'(1,2)[symmetric])
+ hence "?g^n*a'^n dvd ?g^n *b'^n"
+ by (simp only: power_mult_distrib mult_commute)
+ with zn z n have th0:"a'^n dvd b'^n" by auto
+ have "a' dvd a'^n" by (simp add: m)
+ with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
+ hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
+ from nat_coprime_dvd_mult[OF nat_coprime_exp [OF ab'(3), of m]] th1
+ have "a' dvd b'" by (subst (asm) mult_commute, blast)
+ hence "a'*?g dvd b'*?g" by simp
+ with ab'(1,2) have ?thesis by simp }
+ ultimately show ?thesis by blast
+qed
+
+lemma int_pow_divides_pow:
+ assumes ab: "(a::int) ^ n dvd b ^n" and n:"n \<noteq> 0"
+ shows "a dvd b"
proof-
- from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
- with gcd_unique[of "gcd u v" x y] show ?thesis by auto
+ let ?g = "gcd a b"
+ from n obtain m where m: "n = Suc m" by (cases n, simp_all)
+ {assume "?g = 0" with ab n have ?thesis by auto }
+ moreover
+ {assume z: "?g \<noteq> 0"
+ hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
+ from int_gcd_coprime_exists[OF z]
+ obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
+ by blast
+ from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n"
+ by (simp add: ab'(1,2)[symmetric])
+ hence "?g^n*a'^n dvd ?g^n *b'^n"
+ by (simp only: power_mult_distrib mult_commute)
+ with zn z n have th0:"a'^n dvd b'^n" by auto
+ have "a' dvd a'^n" by (simp add: m)
+ with th0 have "a' dvd b'^n"
+ using dvd_trans[of a' "a'^n" "b'^n"] by simp
+ hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
+ from int_coprime_dvd_mult[OF int_coprime_exp [OF ab'(3), of m]] th1
+ have "a' dvd b'" by (subst (asm) mult_commute, blast)
+ hence "a'*?g dvd b'*?g" by simp
+ with ab'(1,2) have ?thesis by simp }
+ ultimately show ?thesis by blast
+qed
+
+lemma nat_pow_divides_eq [simp]: "n ~= 0 \<Longrightarrow> ((a::nat)^n dvd b^n) = (a dvd b)"
+ by (auto intro: nat_pow_divides_pow dvd_power_same)
+
+lemma int_pow_divides_eq [simp]: "n ~= 0 \<Longrightarrow> ((a::int)^n dvd b^n) = (a dvd b)"
+ by (auto intro: int_pow_divides_pow dvd_power_same)
+
+lemma nat_divides_mult:
+ assumes mr: "(m::nat) dvd r" and nr: "n dvd r" and mn:"coprime m n"
+ shows "m * n dvd r"
+proof-
+ from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
+ unfolding dvd_def by blast
+ from mr n' have "m dvd n'*n" by (simp add: mult_commute)
+ hence "m dvd n'" using nat_coprime_dvd_mult_iff[OF mn] by simp
+ then obtain k where k: "n' = m*k" unfolding dvd_def by blast
+ from n' k show ?thesis unfolding dvd_def by auto
+qed
+
+lemma int_divides_mult:
+ assumes mr: "(m::int) dvd r" and nr: "n dvd r" and mn:"coprime m n"
+ shows "m * n dvd r"
+proof-
+ from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
+ unfolding dvd_def by blast
+ from mr n' have "m dvd n'*n" by (simp add: mult_commute)
+ hence "m dvd n'" using int_coprime_dvd_mult_iff[OF mn] by simp
+ then obtain k where k: "n' = m*k" unfolding dvd_def by blast
+ from n' k show ?thesis unfolding dvd_def by auto
qed
-lemma ind_euclid:
- assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
- and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"
+lemma nat_coprime_plus_one [simp]: "coprime ((n::nat) + 1) n"
+ apply (subgoal_tac "gcd (n + 1) n dvd (n + 1 - n)")
+ apply force
+ apply (rule nat_dvd_diff)
+ apply auto
+done
+
+lemma nat_coprime_Suc [simp]: "coprime (Suc n) n"
+ using nat_coprime_plus_one by (simp add: One_nat_def)
+
+lemma int_coprime_plus_one [simp]: "coprime ((n::int) + 1) n"
+ apply (subgoal_tac "gcd (n + 1) n dvd (n + 1 - n)")
+ apply force
+ apply (rule dvd_diff)
+ apply auto
+done
+
+lemma nat_coprime_minus_one: "(n::nat) \<noteq> 0 \<Longrightarrow> coprime (n - 1) n"
+ using nat_coprime_plus_one [of "n - 1"]
+ nat_gcd_commute [of "n - 1" n] by auto
+
+lemma int_coprime_minus_one: "coprime ((n::int) - 1) n"
+ using int_coprime_plus_one [of "n - 1"]
+ int_gcd_commute [of "n - 1" n] by auto
+
+lemma nat_setprod_coprime [rule_format]:
+ "(ALL i: A. coprime (f i) (x::nat)) --> coprime (PROD i:A. f i) x"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto simp add: nat_gcd_mult_cancel)
+done
+
+lemma int_setprod_coprime [rule_format]:
+ "(ALL i: A. coprime (f i) (x::int)) --> coprime (PROD i:A. f i) x"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto simp add: int_gcd_mult_cancel)
+done
+
+lemma nat_prime_odd: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+ unfolding prime_nat_def
+ apply (subst even_mult_two_ex)
+ apply clarify
+ apply (drule_tac x = 2 in spec)
+ apply auto
+done
+
+lemma int_prime_odd: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+ unfolding prime_int_def
+ apply (frule nat_prime_odd)
+ apply (auto simp add: even_nat_def)
+done
+
+lemma nat_coprime_common_divisor: "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow>
+ x dvd b \<Longrightarrow> x = 1"
+ apply (subgoal_tac "x dvd gcd a b")
+ apply simp
+ apply (erule (1) nat_gcd_greatest)
+done
+
+lemma int_coprime_common_divisor: "coprime (a::int) b \<Longrightarrow> x dvd a \<Longrightarrow>
+ x dvd b \<Longrightarrow> abs x = 1"
+ apply (subgoal_tac "x dvd gcd a b")
+ apply simp
+ apply (erule (1) int_gcd_greatest)
+done
+
+lemma nat_coprime_divisors: "(d::int) dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow>
+ coprime d e"
+ apply (auto simp add: dvd_def)
+ apply (frule int_coprime_lmult)
+ apply (subst int_gcd_commute)
+ apply (subst (asm) (2) int_gcd_commute)
+ apply (erule int_coprime_lmult)
+done
+
+lemma nat_invertible_coprime: "(x::nat) * y mod m = 1 \<Longrightarrow> coprime x m"
+apply (metis nat_coprime_lmult nat_gcd_1 nat_gcd_commute nat_gcd_red)
+done
+
+lemma int_invertible_coprime: "(x::int) * y mod m = 1 \<Longrightarrow> coprime x m"
+apply (metis int_coprime_lmult int_gcd_1 int_gcd_commute int_gcd_red)
+done
+
+
+subsection {* Bezout's theorem *}
+
+(* Function bezw returns a pair of witnesses to Bezout's theorem --
+ see the theorems that follow the definition. *)
+fun
+ bezw :: "nat \<Rightarrow> nat \<Rightarrow> int * int"
+where
+ "bezw x y =
+ (if y = 0 then (1, 0) else
+ (snd (bezw y (x mod y)),
+ fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y)))"
+
+lemma bezw_0 [simp]: "bezw x 0 = (1, 0)" by simp
+
+lemma bezw_non_0: "y > 0 \<Longrightarrow> bezw x y = (snd (bezw y (x mod y)),
+ fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y))"
+ by simp
+
+declare bezw.simps [simp del]
+
+lemma bezw_aux [rule_format]:
+ "fst (bezw x y) * int x + snd (bezw x y) * int y = int (gcd x y)"
+proof (induct x y rule: nat_gcd_induct)
+ fix m :: nat
+ show "fst (bezw m 0) * int m + snd (bezw m 0) * int 0 = int (gcd m 0)"
+ by auto
+ next fix m :: nat and n
+ assume ngt0: "n > 0" and
+ ih: "fst (bezw n (m mod n)) * int n +
+ snd (bezw n (m mod n)) * int (m mod n) =
+ int (gcd n (m mod n))"
+ thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
+ apply (simp add: bezw_non_0 nat_gcd_non_0)
+ apply (erule subst)
+ apply (simp add: ring_simps)
+ apply (subst mod_div_equality [of m n, symmetric])
+ (* applying simp here undoes the last substitution!
+ what is procedure cancel_div_mod? *)
+ apply (simp only: ring_simps zadd_int [symmetric]
+ zmult_int [symmetric])
+ done
+qed
+
+lemma int_bezout:
+ fixes x y
+ shows "EX u v. u * (x::int) + v * y = gcd x y"
+proof -
+ have bezout_aux: "!!x y. x \<ge> (0::int) \<Longrightarrow> y \<ge> 0 \<Longrightarrow>
+ EX u v. u * x + v * y = gcd x y"
+ apply (rule_tac x = "fst (bezw (nat x) (nat y))" in exI)
+ apply (rule_tac x = "snd (bezw (nat x) (nat y))" in exI)
+ apply (unfold gcd_int_def)
+ apply simp
+ apply (subst bezw_aux [symmetric])
+ apply auto
+ done
+ have "(x \<ge> 0 \<and> y \<ge> 0) | (x \<ge> 0 \<and> y \<le> 0) | (x \<le> 0 \<and> y \<ge> 0) |
+ (x \<le> 0 \<and> y \<le> 0)"
+ by auto
+ moreover have "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> ?thesis"
+ by (erule (1) bezout_aux)
+ moreover have "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> ?thesis"
+ apply (insert bezout_aux [of x "-y"])
+ apply auto
+ apply (rule_tac x = u in exI)
+ apply (rule_tac x = "-v" in exI)
+ apply (subst int_gcd_neg2 [symmetric])
+ apply auto
+ done
+ moreover have "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> ?thesis"
+ apply (insert bezout_aux [of "-x" y])
+ apply auto
+ apply (rule_tac x = "-u" in exI)
+ apply (rule_tac x = v in exI)
+ apply (subst int_gcd_neg1 [symmetric])
+ apply auto
+ done
+ moreover have "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> ?thesis"
+ apply (insert bezout_aux [of "-x" "-y"])
+ apply auto
+ apply (rule_tac x = "-u" in exI)
+ apply (rule_tac x = "-v" in exI)
+ apply (subst int_gcd_neg1 [symmetric])
+ apply (subst int_gcd_neg2 [symmetric])
+ apply auto
+ done
+ ultimately show ?thesis by blast
+qed
+
+text {* versions of Bezout for nat, by Amine Chaieb *}
+
+lemma ind_euclid:
+ assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
+ and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"
shows "P a b"
proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
fix n a b
assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
have "a = b \<or> a < b \<or> b < a" by arith
moreover {assume eq: "a= b"
- from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
+ from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq
+ by simp}
moreover
{assume lt: "a < b"
hence "a + b - a < n \<or> a = 0" using H(2) by arith
@@ -269,65 +1180,67 @@
ultimately show "P a b" by blast
qed
-lemma bezout_lemma:
- assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
- shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
-using ex
-apply clarsimp
-apply (rule_tac x="d" in exI, simp add: dvd_add)
-apply (case_tac "a * x = b * y + d" , simp_all)
-apply (rule_tac x="x + y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x + y" in exI)
-apply algebra
+lemma nat_bezout_lemma:
+ assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
+ (a * x = b * y + d \<or> b * x = a * y + d)"
+ shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and>
+ (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
+ using ex
+ apply clarsimp
+ apply (rule_tac x="d" in exI, simp add: dvd_add)
+ apply (case_tac "a * x = b * y + d" , simp_all)
+ apply (rule_tac x="x + y" in exI)
+ apply (rule_tac x="y" in exI)
+ apply algebra
+ apply (rule_tac x="x" in exI)
+ apply (rule_tac x="x + y" in exI)
+ apply algebra
done
-lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
-apply(induct a b rule: ind_euclid)
-apply blast
-apply clarify
-apply (rule_tac x="a" in exI, simp add: dvd_add)
-apply clarsimp
-apply (rule_tac x="d" in exI)
-apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
-apply (rule_tac x="x+y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x+y" in exI)
-apply algebra
+lemma nat_bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
+ (a * x = b * y + d \<or> b * x = a * y + d)"
+ apply(induct a b rule: ind_euclid)
+ apply blast
+ apply clarify
+ apply (rule_tac x="a" in exI, simp add: dvd_add)
+ apply clarsimp
+ apply (rule_tac x="d" in exI)
+ apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
+ apply (rule_tac x="x+y" in exI)
+ apply (rule_tac x="y" in exI)
+ apply algebra
+ apply (rule_tac x="x" in exI)
+ apply (rule_tac x="x+y" in exI)
+ apply algebra
done
-lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
-using bezout_add[of a b]
-apply clarsimp
-apply (rule_tac x="d" in exI, simp)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="y" in exI)
-apply auto
+lemma nat_bezout1: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
+ (a * x - b * y = d \<or> b * x - a * y = d)"
+ using nat_bezout_add[of a b]
+ apply clarsimp
+ apply (rule_tac x="d" in exI, simp)
+ apply (rule_tac x="x" in exI)
+ apply (rule_tac x="y" in exI)
+ apply auto
done
-
-text {* We can get a stronger version with a nonzeroness assumption. *}
-lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
-
-lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
+lemma nat_bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
proof-
- from nz have ap: "a > 0" by simp
- from bezout_add[of a b]
- have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
+ from nz have ap: "a > 0" by simp
+ from nat_bezout_add[of a b]
+ have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or>
+ (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
moreover
- {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
- from H have ?thesis by blast }
+ {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
+ from H have ?thesis by blast }
moreover
{fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
{assume b0: "b = 0" with H have ?thesis by simp}
- moreover
+ moreover
{assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
- from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
+ from b dvd_imp_le [OF H(2)] have "d < b \<or> d = b"
+ by auto
moreover
{assume db: "d=b"
from prems have ?thesis apply simp
@@ -335,18 +1248,22 @@
apply (rule exI[where x = b])
by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
moreover
- {assume db: "d < b"
+ {assume db: "d < b"
{assume "x=0" hence ?thesis using prems by simp }
moreover
{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
-
from db have "d \<le> b - 1" by simp
hence "d*b \<le> b*(b - 1)" by simp
with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
- from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
+ from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)"
+ by simp
+ hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"
+ by (simp only: mult_assoc right_distrib)
+ hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)"
+ by algebra
hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
- hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
+ hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
by (simp only: diff_add_assoc[OF dble, of d, symmetric])
hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
by (simp only: diff_mult_distrib2 add_commute mult_ac)
@@ -361,156 +1278,155 @@
ultimately show ?thesis by blast
qed
-
-lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
-proof-
- let ?g = "gcd a b"
- from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
- from d(1,2) have "d dvd ?g" by simp
- then obtain k where k: "?g = d*k" unfolding dvd_def by blast
- from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast
- hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k"
- by (algebra add: diff_mult_distrib)
- hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g"
- by (simp add: k mult_assoc)
- thus ?thesis by blast
-qed
-
-lemma bezout_gcd_strong: assumes a: "a \<noteq> 0"
+lemma nat_bezout: assumes a: "(a::nat) \<noteq> 0"
shows "\<exists>x y. a * x = b * y + gcd a b"
proof-
let ?g = "gcd a b"
- from bezout_add_strong[OF a, of b]
+ from nat_bezout_add_strong[OF a, of b]
obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
from d(1,2) have "d dvd ?g" by simp
then obtain k where k: "?g = d*k" unfolding dvd_def by blast
- from d(3) have "a * x * k = (b * y + d) *k " by algebra
+ from d(3) have "a * x * k = (b * y + d) *k " by auto
hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
thus ?thesis by blast
qed
-lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
-by(simp add: gcd_mult_distrib2 mult_commute)
+
+subsection {* LCM *}
+
+lemma int_lcm_altdef: "lcm (a::int) b = (abs a) * (abs b) div gcd a b"
+ by (simp add: lcm_int_def lcm_nat_def zdiv_int
+ zmult_int [symmetric] gcd_int_def)
+
+lemma nat_prod_gcd_lcm: "(m::nat) * n = gcd m n * lcm m n"
+ unfolding lcm_nat_def
+ by (simp add: dvd_mult_div_cancel [OF nat_gcd_dvd_prod])
+
+lemma int_prod_gcd_lcm: "abs(m::int) * abs n = gcd m n * lcm m n"
+ unfolding lcm_int_def gcd_int_def
+ apply (subst int_mult [symmetric])
+ apply (subst nat_prod_gcd_lcm [symmetric])
+ apply (subst nat_abs_mult_distrib [symmetric])
+ apply (simp, simp add: abs_mult)
+done
+
+lemma nat_lcm_0 [simp]: "lcm (m::nat) 0 = 0"
+ unfolding lcm_nat_def by simp
+
+lemma int_lcm_0 [simp]: "lcm (m::int) 0 = 0"
+ unfolding lcm_int_def by simp
+
+lemma nat_lcm_1 [simp]: "lcm (m::nat) 1 = m"
+ unfolding lcm_nat_def by simp
+
+lemma nat_lcm_Suc_0 [simp]: "lcm (m::nat) (Suc 0) = m"
+ unfolding lcm_nat_def by (simp add: One_nat_def)
+
+lemma int_lcm_1 [simp]: "lcm (m::int) 1 = abs m"
+ unfolding lcm_int_def by simp
+
+lemma nat_lcm_0_left [simp]: "lcm (0::nat) n = 0"
+ unfolding lcm_nat_def by simp
-lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
- (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
- let ?g = "gcd a b"
- {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
- from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
- by blast
- hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
- hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k"
- by (simp only: diff_mult_distrib)
- hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
- by (simp add: k[symmetric] mult_assoc)
- hence ?lhs by blast}
+lemma int_lcm_0_left [simp]: "lcm (0::int) n = 0"
+ unfolding lcm_int_def by simp
+
+lemma nat_lcm_1_left [simp]: "lcm (1::nat) m = m"
+ unfolding lcm_nat_def by simp
+
+lemma nat_lcm_Suc_0_left [simp]: "lcm (Suc 0) m = m"
+ unfolding lcm_nat_def by (simp add: One_nat_def)
+
+lemma int_lcm_1_left [simp]: "lcm (1::int) m = abs m"
+ unfolding lcm_int_def by simp
+
+lemma nat_lcm_commute: "lcm (m::nat) n = lcm n m"
+ unfolding lcm_nat_def by (simp add: nat_gcd_commute ring_simps)
+
+lemma int_lcm_commute: "lcm (m::int) n = lcm n m"
+ unfolding lcm_int_def by (subst nat_lcm_commute, rule refl)
+
+
+lemma nat_lcm_pos:
+ assumes mpos: "(m::nat) > 0"
+ and npos: "n>0"
+ shows "lcm m n > 0"
+proof(rule ccontr, simp add: lcm_nat_def nat_gcd_zero)
+ assume h:"m*n div gcd m n = 0"
+ from mpos npos have "gcd m n \<noteq> 0" using nat_gcd_zero by simp
+ hence gcdp: "gcd m n > 0" by simp
+ with h
+ have "m*n < gcd m n"
+ by (cases "m * n < gcd m n")
+ (auto simp add: div_if[OF gcdp, where m="m*n"])
moreover
- {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
- have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
- using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
- from nat_dvd_diff[OF dv(1,2)] nat_dvd_diff[OF dv(3,4)] H
- have ?rhs by auto}
- ultimately show ?thesis by blast
-qed
-
-lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
-proof-
- let ?g = "gcd a b"
- have dv: "?g dvd a*x" "?g dvd b * y"
- using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
- from dvd_add[OF dv] H
- show ?thesis by auto
+ have "gcd m n dvd m" by simp
+ with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
+ with npos have t1:"gcd m n*n \<le> m*n" by simp
+ have "gcd m n \<le> gcd m n*n" using npos by simp
+ with t1 have "gcd m n \<le> m*n" by arith
+ ultimately show "False" by simp
qed
-lemma gcd_mult': "gcd b (a * b) = b"
-by (simp add: gcd_mult mult_commute[of a b])
-
-lemma gcd_add: "gcd(a + b) b = gcd a b"
- "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
-apply (simp_all add: gcd_add1)
-by (simp add: gcd_commute gcd_add1)
-
-lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
-proof-
- {fix a b assume H: "b \<le> (a::nat)"
- hence th: "a - b + b = a" by arith
- from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp}
- note th = this
-{
- assume ab: "b \<le> a"
- from th[OF ab] show "gcd (a - b) b = gcd a b" by blast
-next
- assume ab: "a \<le> b"
- from th[OF ab] show "gcd a (b - a) = gcd a b"
- by (simp add: gcd_commute)}
-qed
-
+lemma int_lcm_pos:
+ assumes mneq0: "(m::int) ~= 0"
+ and npos: "n ~= 0"
+ shows "lcm m n > 0"
-subsection {* LCM defined by GCD *}
-
-
-definition
- lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
- lcm_def: "lcm m n = m * n div gcd m n"
-
-lemma prod_gcd_lcm:
- "m * n = gcd m n * lcm m n"
- unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
+ apply (subst int_lcm_abs)
+ apply (rule nat_lcm_pos [transferred])
+ using prems apply auto
+done
-lemma lcm_0 [simp]: "lcm m 0 = 0"
- unfolding lcm_def by simp
-
-lemma lcm_1 [simp]: "lcm m 1 = m"
- unfolding lcm_def by simp
-
-lemma lcm_0_left [simp]: "lcm 0 n = 0"
- unfolding lcm_def by simp
-
-lemma lcm_1_left [simp]: "lcm 1 m = m"
- unfolding lcm_def by simp
-
-lemma dvd_pos:
+lemma nat_dvd_pos:
fixes n m :: nat
assumes "n > 0" and "m dvd n"
shows "m > 0"
using assms by (cases m) auto
-lemma lcm_least:
- assumes "m dvd k" and "n dvd k"
+lemma nat_lcm_least:
+ assumes "(m::nat) dvd k" and "n dvd k"
shows "lcm m n dvd k"
proof (cases k)
case 0 then show ?thesis by auto
next
case (Suc _) then have pos_k: "k > 0" by auto
- from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
- with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
+ from assms nat_dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
+ with nat_gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
from assms obtain p where k_m: "k = m * p" using dvd_def by blast
from assms obtain q where k_n: "k = n * q" using dvd_def by blast
from pos_k k_m have pos_p: "p > 0" by auto
from pos_k k_n have pos_q: "q > 0" by auto
have "k * k * gcd q p = k * gcd (k * q) (k * p)"
- by (simp add: mult_ac gcd_mult_distrib2)
+ by (simp add: mult_ac nat_gcd_mult_distrib)
also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
by (simp add: k_m [symmetric] k_n [symmetric])
also have "\<dots> = k * p * q * gcd m n"
- by (simp add: mult_ac gcd_mult_distrib2)
+ by (simp add: mult_ac nat_gcd_mult_distrib)
finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
by (simp only: k_m [symmetric] k_n [symmetric])
then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
by (simp add: mult_ac)
with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
by simp
- with prod_gcd_lcm [of m n]
+ with nat_prod_gcd_lcm [of m n]
have "lcm m n * gcd q p * gcd m n = k * gcd m n"
by (simp add: mult_ac)
- with pos_gcd have "lcm m n * gcd q p = k" by simp
+ with pos_gcd have "lcm m n * gcd q p = k" by auto
then show ?thesis using dvd_def by auto
qed
-lemma lcm_dvd1 [iff]:
- "m dvd lcm m n"
+lemma int_lcm_least:
+ assumes "(m::int) dvd k" and "n dvd k"
+ shows "lcm m n dvd k"
+
+ apply (subst int_lcm_abs)
+ apply (rule dvd_trans)
+ apply (rule nat_lcm_least [transferred, of _ "abs k" _])
+ using prems apply auto
+done
+
+lemma nat_lcm_dvd1: "(m::nat) dvd lcm m n"
proof (cases m)
case 0 then show ?thesis by simp
next
@@ -524,264 +1440,354 @@
then have npos: "n > 0" by simp
have "gcd m n dvd n" by simp
then obtain k where "n = gcd m n * k" using dvd_def by auto
- then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
- also have "\<dots> = m * k" using mpos npos gcd_zero by simp
- finally show ?thesis by (simp add: lcm_def)
- qed
-qed
-
-lemma lcm_dvd2 [iff]:
- "n dvd lcm m n"
-proof (cases n)
- case 0 then show ?thesis by simp
-next
- case (Suc _)
- then have npos: "n > 0" by simp
- show ?thesis
- proof (cases m)
- case 0 then show ?thesis by simp
- next
- case (Suc _)
- then have mpos: "m > 0" by simp
- have "gcd m n dvd m" by simp
- then obtain k where "m = gcd m n * k" using dvd_def by auto
- then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
- also have "\<dots> = n * k" using mpos npos gcd_zero by simp
- finally show ?thesis by (simp add: lcm_def)
+ then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n"
+ by (simp add: mult_ac)
+ also have "\<dots> = m * k" using mpos npos nat_gcd_zero by simp
+ finally show ?thesis by (simp add: lcm_nat_def)
qed
qed
-lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
- by (simp add: gcd_commute)
+lemma int_lcm_dvd1: "(m::int) dvd lcm m n"
+ apply (subst int_lcm_abs)
+ apply (rule dvd_trans)
+ prefer 2
+ apply (rule nat_lcm_dvd1 [transferred])
+ apply auto
+done
+
+lemma nat_lcm_dvd2: "(n::nat) dvd lcm m n"
+ by (subst nat_lcm_commute, rule nat_lcm_dvd1)
+
+lemma int_lcm_dvd2: "(n::int) dvd lcm m n"
+ by (subst int_lcm_commute, rule int_lcm_dvd1)
+
+lemma dvd_lcm_I1_nat[simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
+by(metis nat_lcm_dvd1 dvd_trans)
+
+lemma dvd_lcm_I2_nat[simp]: "(k::nat) dvd n \<Longrightarrow> k dvd lcm m n"
+by(metis nat_lcm_dvd2 dvd_trans)
+
+lemma dvd_lcm_I1_int[simp]: "(i::int) dvd m \<Longrightarrow> i dvd lcm m n"
+by(metis int_lcm_dvd1 dvd_trans)
+
+lemma dvd_lcm_I2_int[simp]: "(i::int) dvd n \<Longrightarrow> i dvd lcm m n"
+by(metis int_lcm_dvd2 dvd_trans)
+
+lemma nat_lcm_unique: "(a::nat) dvd d \<and> b dvd d \<and>
+ (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
+ by (auto intro: dvd_anti_sym nat_lcm_least nat_lcm_dvd1 nat_lcm_dvd2)
+
+lemma int_lcm_unique: "d >= 0 \<and> (a::int) dvd d \<and> b dvd d \<and>
+ (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
+ by (auto intro: dvd_anti_sym [transferred] int_lcm_least)
-lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
- apply (subgoal_tac "n = m + (n - m)")
- apply (erule ssubst, rule gcd_add1_eq, simp)
- done
+lemma nat_lcm_dvd_eq [simp]: "(x::nat) dvd y \<Longrightarrow> lcm x y = y"
+ apply (rule sym)
+ apply (subst nat_lcm_unique [symmetric])
+ apply auto
+done
+
+lemma int_lcm_dvd_eq [simp]: "0 <= y \<Longrightarrow> (x::int) dvd y \<Longrightarrow> lcm x y = y"
+ apply (rule sym)
+ apply (subst int_lcm_unique [symmetric])
+ apply auto
+done
+
+lemma nat_lcm_dvd_eq' [simp]: "(x::nat) dvd y \<Longrightarrow> lcm y x = y"
+ by (subst nat_lcm_commute, erule nat_lcm_dvd_eq)
+
+lemma int_lcm_dvd_eq' [simp]: "y >= 0 \<Longrightarrow> (x::int) dvd y \<Longrightarrow> lcm y x = y"
+ by (subst int_lcm_commute, erule (1) int_lcm_dvd_eq)
+
+
+lemma lcm_assoc_nat: "lcm (lcm n m) (p::nat) = lcm n (lcm m p)"
+apply(rule nat_lcm_unique[THEN iffD1])
+apply (metis dvd.order_trans nat_lcm_unique)
+done
+
+lemma lcm_assoc_int: "lcm (lcm n m) (p::int) = lcm n (lcm m p)"
+apply(rule int_lcm_unique[THEN iffD1])
+apply (metis dvd_trans int_lcm_unique)
+done
+
+lemmas lcm_left_commute_nat =
+ mk_left_commute[of lcm, OF lcm_assoc_nat nat_lcm_commute]
+
+lemmas lcm_left_commute_int =
+ mk_left_commute[of lcm, OF lcm_assoc_int int_lcm_commute]
+
+lemmas lcm_ac_nat = lcm_assoc_nat nat_lcm_commute lcm_left_commute_nat
+lemmas lcm_ac_int = lcm_assoc_int int_lcm_commute lcm_left_commute_int
-subsection {* GCD and LCM on integers *}
-
-definition
- zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
- "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
+subsection {* Primes *}
-lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i"
-by (simp add: zgcd_def int_dvd_iff)
+(* Is there a better way to handle these, rather than making them
+ elim rules? *)
-lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j"
-by (simp add: zgcd_def int_dvd_iff)
+lemma nat_prime_ge_0 [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
+ by (unfold prime_nat_def, auto)
-lemma zgcd_pos: "zgcd i j \<ge> 0"
-by (simp add: zgcd_def)
+lemma nat_prime_gt_0 [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
+ by (unfold prime_nat_def, auto)
-lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
-by (simp add: zgcd_def gcd_zero)
+lemma nat_prime_ge_1 [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
+ by (unfold prime_nat_def, auto)
-lemma zgcd_commute: "zgcd i j = zgcd j i"
-unfolding zgcd_def by (simp add: gcd_commute)
+lemma nat_prime_gt_1 [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
+ by (unfold prime_nat_def, auto)
-lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
-unfolding zgcd_def by simp
+lemma nat_prime_ge_Suc_0 [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
+ by (unfold prime_nat_def, auto)
-lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
-unfolding zgcd_def by simp
+lemma nat_prime_gt_Suc_0 [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
+ by (unfold prime_nat_def, auto)
+
+lemma nat_prime_ge_2 [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
+ by (unfold prime_nat_def, auto)
+
+lemma int_prime_ge_0 [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
+ by (unfold prime_int_def prime_nat_def, auto)
- (* should be solved by algebra*)
-lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
- unfolding zgcd_def
-proof -
- assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
- then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
- from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
- have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
- unfolding dvd_def
- by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
- from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
- unfolding dvd_def by blast
- from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
- then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
- then show ?thesis
- apply (subst abs_dvd_iff [symmetric])
- apply (subst dvd_abs_iff [symmetric])
- apply (unfold dvd_def)
- apply (rule_tac x = "int h'" in exI, simp)
- done
-qed
+lemma int_prime_gt_0 [elim]: "prime (p::int) \<Longrightarrow> p > 0"
+ by (unfold prime_int_def prime_nat_def, auto)
+
+lemma int_prime_ge_1 [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
+ by (unfold prime_int_def prime_nat_def, auto)
-lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
+lemma int_prime_gt_1 [elim]: "prime (p::int) \<Longrightarrow> p > 1"
+ by (unfold prime_int_def prime_nat_def, auto)
+
+lemma int_prime_ge_2 [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
+ by (unfold prime_int_def prime_nat_def, auto)
-lemma zgcd_greatest:
- assumes "k dvd m" and "k dvd n"
- shows "k dvd zgcd m n"
-proof -
- let ?k' = "nat \<bar>k\<bar>"
- let ?m' = "nat \<bar>m\<bar>"
- let ?n' = "nat \<bar>n\<bar>"
- from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
- unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
- from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
- unfolding zgcd_def by (simp only: zdvd_int)
- then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
- then show "k dvd zgcd m n" by simp
-qed
+thm prime_nat_def;
+thm prime_nat_def [transferred];
+
+lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow>
+ m = 1 \<or> m = p))"
+ using prime_nat_def [transferred]
+ apply (case_tac "p >= 0")
+ by (blast, auto simp add: int_prime_ge_0)
+
+(* To do: determine primality of any numeral *)
+
+lemma nat_zero_not_prime [simp]: "~prime (0::nat)"
+ by (simp add: prime_nat_def)
+
+lemma int_zero_not_prime [simp]: "~prime (0::int)"
+ by (simp add: prime_int_def)
+
+lemma nat_one_not_prime [simp]: "~prime (1::nat)"
+ by (simp add: prime_nat_def)
-lemma div_zgcd_relprime:
- assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
- shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
-proof -
- from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith
- let ?g = "zgcd a b"
- let ?a' = "a div ?g"
- let ?b' = "b div ?g"
- let ?g' = "zgcd ?a' ?b'"
- have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
- have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
- from dvdg dvdg' obtain ka kb ka' kb' where
- kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
- unfolding dvd_def by blast
- then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
- then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
- by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
- zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
- have "?g \<noteq> 0" using nz by simp
- then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
- from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
- with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
- with zgcd_pos show "?g' = 1" by simp
-qed
+lemma nat_Suc_0_not_prime [simp]: "~prime (Suc 0)"
+ by (simp add: prime_nat_def One_nat_def)
+
+lemma int_one_not_prime [simp]: "~prime (1::int)"
+ by (simp add: prime_int_def)
+
+lemma nat_two_is_prime [simp]: "prime (2::nat)"
+ apply (auto simp add: prime_nat_def)
+ apply (case_tac m)
+ apply (auto dest!: dvd_imp_le)
+ done
-lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
- by (simp add: zgcd_def abs_if)
+lemma int_two_is_prime [simp]: "prime (2::int)"
+ by (rule nat_two_is_prime [transferred direction: nat "op <= (0::int)"])
-lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
- apply (frule_tac b = n and a = m in pos_mod_sign)
- apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
- apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
- apply (frule_tac a = m in pos_mod_bound)
- apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
+lemma nat_prime_imp_coprime: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+ apply (unfold prime_nat_def)
+ apply (metis nat_gcd_dvd1 nat_gcd_dvd2)
+ done
+
+lemma int_prime_imp_coprime: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+ apply (unfold prime_int_altdef)
+ apply (metis int_gcd_dvd1 int_gcd_dvd2 int_gcd_ge_0)
done
-lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
- apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
- apply (auto simp add: linorder_neq_iff zgcd_non_0)
- apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
- done
+lemma nat_prime_dvd_mult: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+ by (blast intro: nat_coprime_dvd_mult nat_prime_imp_coprime)
+
+lemma int_prime_dvd_mult: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+ by (blast intro: int_coprime_dvd_mult int_prime_imp_coprime)
+
+lemma nat_prime_dvd_mult_eq [simp]: "prime (p::nat) \<Longrightarrow>
+ p dvd m * n = (p dvd m \<or> p dvd n)"
+ by (rule iffI, rule nat_prime_dvd_mult, auto)
-lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
- by (simp add: zgcd_def abs_if)
+lemma int_prime_dvd_mult_eq [simp]: "prime (p::int) \<Longrightarrow>
+ p dvd m * n = (p dvd m \<or> p dvd n)"
+ by (rule iffI, rule int_prime_dvd_mult, auto)
-lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
- by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
+lemma nat_not_prime_eq_prod: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+ EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
+ unfolding prime_nat_def dvd_def apply auto
+ apply (subgoal_tac "k > 1")
+ apply force
+ apply (subgoal_tac "k ~= 0")
+ apply force
+ apply (rule notI)
+ apply force
+done
-lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
- by (simp add: zgcd_def gcd_1_left)
-
-lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
- by (simp add: zgcd_def gcd_assoc)
+(* there's a lot of messing around with signs of products here --
+ could this be made more automatic? *)
+lemma int_not_prime_eq_prod: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+ EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
+ unfolding prime_int_altdef dvd_def
+ apply auto
+ apply (rule_tac x = m in exI)
+ apply (rule_tac x = k in exI)
+ apply (auto simp add: mult_compare_simps)
+ apply (subgoal_tac "k > 0")
+ apply arith
+ apply (case_tac "k <= 0")
+ apply (subgoal_tac "m * k <= 0")
+ apply force
+ apply (subst zero_compare_simps(8))
+ apply auto
+ apply (subgoal_tac "m * k <= 0")
+ apply force
+ apply (subst zero_compare_simps(8))
+ apply auto
+done
-lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
- apply (rule zgcd_commute [THEN trans])
- apply (rule zgcd_assoc [THEN trans])
- apply (rule zgcd_commute [THEN arg_cong])
- done
+lemma nat_prime_dvd_power [rule_format]: "prime (p::nat) -->
+ n > 0 --> (p dvd x^n --> p dvd x)"
+ by (induct n rule: nat_induct, auto)
-lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
- -- {* addition is an AC-operator *}
+lemma int_prime_dvd_power [rule_format]: "prime (p::int) -->
+ n > 0 --> (p dvd x^n --> p dvd x)"
+ apply (induct n rule: nat_induct, auto)
+ apply (frule int_prime_ge_0)
+ apply auto
+done
+
+lemma nat_prime_imp_power_coprime: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
+ coprime a (p^m)"
+ apply (rule nat_coprime_exp)
+ apply (subst nat_gcd_commute)
+ apply (erule (1) nat_prime_imp_coprime)
+done
-lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
- by (simp del: minus_mult_right [symmetric]
- add: minus_mult_right nat_mult_distrib zgcd_def abs_if
- mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
+lemma int_prime_imp_power_coprime: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
+ coprime a (p^m)"
+ apply (rule int_coprime_exp)
+ apply (subst int_gcd_commute)
+ apply (erule (1) int_prime_imp_coprime)
+done
-lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
- by (simp add: abs_if zgcd_zmult_distrib2)
+lemma nat_primes_coprime: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+ apply (rule nat_prime_imp_coprime, assumption)
+ apply (unfold prime_nat_def, auto)
+done
-lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
- by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
+lemma int_primes_coprime: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+ apply (rule int_prime_imp_coprime, assumption)
+ apply (unfold prime_int_altdef, clarify)
+ apply (drule_tac x = q in spec)
+ apply (drule_tac x = p in spec)
+ apply auto
+done
-lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
- by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
+lemma nat_primes_imp_powers_coprime: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
+ coprime (p^m) (q^n)"
+ by (rule nat_coprime_exp2, rule nat_primes_coprime)
-lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
- by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
+lemma int_primes_imp_powers_coprime: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
+ coprime (p^m) (q^n)"
+ by (rule int_coprime_exp2, rule int_primes_coprime)
-
-definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
+lemma nat_prime_factor: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
+ apply (induct n rule: nat_less_induct)
+ apply (case_tac "n = 0")
+ using nat_two_is_prime apply blast
+ apply (case_tac "prime n")
+ apply blast
+ apply (subgoal_tac "n > 1")
+ apply (frule (1) nat_not_prime_eq_prod)
+ apply (auto intro: dvd_mult dvd_mult2)
+done
-lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
+(* An Isar version:
+
+lemma nat_prime_factor_b:
+ fixes n :: nat
+ assumes "n \<noteq> 1"
+ shows "\<exists>p. prime p \<and> p dvd n"
-lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
-
-
-lemma dvd_imp_dvd_zlcm1:
- assumes "k dvd i" shows "k dvd (zlcm i j)"
-proof -
- have "nat(abs k) dvd nat(abs i)" using `k dvd i`
- by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
- thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+using `n ~= 1`
+proof (induct n rule: nat_less_induct)
+ fix n :: nat
+ assume "n ~= 1" and
+ ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)"
+ thus "\<exists>p. prime p \<and> p dvd n"
+ proof -
+ {
+ assume "n = 0"
+ moreover note nat_two_is_prime
+ ultimately have ?thesis
+ by (auto simp del: nat_two_is_prime)
+ }
+ moreover
+ {
+ assume "prime n"
+ hence ?thesis by auto
+ }
+ moreover
+ {
+ assume "n ~= 0" and "~ prime n"
+ with `n ~= 1` have "n > 1" by auto
+ with `~ prime n` and nat_not_prime_eq_prod obtain m k where
+ "n = m * k" and "1 < m" and "m < n" by blast
+ with ih obtain p where "prime p" and "p dvd m" by blast
+ with `n = m * k` have ?thesis by auto
+ }
+ ultimately show ?thesis by blast
+ qed
qed
-lemma dvd_imp_dvd_zlcm2:
- assumes "k dvd j" shows "k dvd (zlcm i j)"
-proof -
- have "nat(abs k) dvd nat(abs j)" using `k dvd j`
- by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
- thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+*)
+
+text {* One property of coprimality is easier to prove via prime factors. *}
+
+lemma nat_prime_divprod_pow:
+ assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b"
+ shows "p^n dvd a \<or> p^n dvd b"
+proof-
+ {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
+ apply (cases "n=0", simp_all)
+ apply (cases "a=1", simp_all) done}
+ moreover
+ {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
+ then obtain m where m: "n = Suc m" by (cases n, auto)
+ from n have "p dvd p^n" by (intro dvd_power, auto)
+ also note pab
+ finally have pab': "p dvd a * b".
+ from nat_prime_dvd_mult[OF p pab']
+ have "p dvd a \<or> p dvd b" .
+ moreover
+ {assume pa: "p dvd a"
+ have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+ from nat_coprime_common_divisor [OF ab, OF pa] p have "\<not> p dvd b" by auto
+ with p have "coprime b p"
+ by (subst nat_gcd_commute, intro nat_prime_imp_coprime)
+ hence pnb: "coprime (p^n) b"
+ by (subst nat_gcd_commute, rule nat_coprime_exp)
+ from nat_coprime_divprod[OF pnba pnb] have ?thesis by blast }
+ moreover
+ {assume pb: "p dvd b"
+ have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+ from nat_coprime_common_divisor [OF ab, of p] pb p have "\<not> p dvd a"
+ by auto
+ with p have "coprime a p"
+ by (subst nat_gcd_commute, intro nat_prime_imp_coprime)
+ hence pna: "coprime (p^n) a"
+ by (subst nat_gcd_commute, rule nat_coprime_exp)
+ from nat_coprime_divprod[OF pab pna] have ?thesis by blast }
+ ultimately have ?thesis by blast}
+ ultimately show ?thesis by blast
qed
-
-lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
-by (case_tac "d <0", simp_all)
-
-lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
-by (case_tac "d<0", simp_all)
-
-(* lcm a b is positive for positive a and b *)
-
-lemma lcm_pos:
- assumes mpos: "m > 0"
- and npos: "n>0"
- shows "lcm m n > 0"
-proof(rule ccontr, simp add: lcm_def gcd_zero)
-assume h:"m*n div gcd m n = 0"
-from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
-hence gcdp: "gcd m n > 0" by simp
-with h
-have "m*n < gcd m n"
- by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
-moreover
-have "gcd m n dvd m" by simp
- with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
- with npos have t1:"gcd m n *n \<le> m*n" by simp
- have "gcd m n \<le> gcd m n*n" using npos by simp
- with t1 have "gcd m n \<le> m*n" by arith
-ultimately show "False" by simp
-qed
-
-lemma zlcm_pos:
- assumes anz: "a \<noteq> 0"
- and bnz: "b \<noteq> 0"
- shows "0 < zlcm a b"
-proof-
- let ?na = "nat (abs a)"
- let ?nb = "nat (abs b)"
- have nap: "?na >0" using anz by simp
- have nbp: "?nb >0" using bnz by simp
- have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
- thus ?thesis by (simp add: zlcm_def)
-qed
-
-lemma zgcd_code [code]:
- "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
- by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
-
end
--- a/src/HOL/Hilbert_Choice.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Hilbert_Choice.thy Tue Jun 23 15:48:55 2009 +0100
@@ -7,7 +7,7 @@
theory Hilbert_Choice
imports Nat Wellfounded Plain
-uses ("Tools/meson.ML") ("Tools/specification_package.ML")
+uses ("Tools/meson.ML") ("Tools/choice_specification.ML")
begin
subsection {* Hilbert's epsilon *}
@@ -596,7 +596,7 @@
lemma exE_some: "[| Ex P ; c == Eps P |] ==> P c"
by (simp only: someI_ex)
-use "Tools/specification_package.ML"
+use "Tools/choice_specification.ML"
end
--- a/src/HOL/HoareParallel/OG_Syntax.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/HoareParallel/OG_Syntax.thy Tue Jun 23 15:48:55 2009 +0100
@@ -95,7 +95,7 @@
| annbexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/HoareParallel/RG_Syntax.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/HoareParallel/RG_Syntax.thy Tue Jun 23 15:48:55 2009 +0100
@@ -68,7 +68,7 @@
| bexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Tue Jun 23 15:48:55 2009 +0100
@@ -318,7 +318,7 @@
val dummy_case_term = IVar dummy_name;
(*assumption: dummy values are not relevant for serialization*)
val unitt = IConst (unit', (([], []), []));
- fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t)
+ fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)
| dest_abs (t, ty) =
let
val vs = Code_Thingol.fold_varnames cons t [];
@@ -337,7 +337,7 @@
then tr_bind' [(x1, ty1), (x2, ty2)]
else force t
| _ => force t;
- in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type),
+ in (dummy_name, dummy_type) `|=> ICase (((IVar dummy_name, dummy_type),
[(unitt, tr_bind' ts)]), dummy_case_term) end
and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys)
of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)]
@@ -349,7 +349,7 @@
| imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t
of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts
| (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts)
- | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t
+ | imp_monad_bind bind' return' unit' (v_ty `|=> t) = v_ty `|=> imp_monad_bind bind' return' unit' t
| imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase
(((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0);
--- a/src/HOL/Import/HOL4Setup.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Import/HOL4Setup.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,10 +1,9 @@
(* Title: HOL/Import/HOL4Setup.thy
- ID: $Id$
Author: Sebastian Skalberg (TU Muenchen)
*)
theory HOL4Setup imports MakeEqual ImportRecorder
- uses ("proof_kernel.ML") ("replay.ML") ("hol4rews.ML") ("import_package.ML") begin
+ uses ("proof_kernel.ML") ("replay.ML") ("hol4rews.ML") ("import.ML") begin
section {* General Setup *}
@@ -162,8 +161,8 @@
use "proof_kernel.ML"
use "replay.ML"
-use "import_package.ML"
+use "import.ML"
-setup ImportPackage.setup
+setup Import.setup
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,71 @@
+(* Title: HOL/Import/import.ML
+ Author: Sebastian Skalberg (TU Muenchen)
+*)
+
+signature IMPORT =
+sig
+ val debug : bool ref
+ val import_tac : Proof.context -> string * string -> tactic
+ val setup : theory -> theory
+end
+
+structure ImportData = TheoryDataFun
+(
+ type T = ProofKernel.thm option array option
+ val empty = NONE
+ val copy = I
+ val extend = I
+ fun merge _ _ = NONE
+)
+
+structure Import :> IMPORT =
+struct
+
+val debug = ref false
+fun message s = if !debug then writeln s else ()
+
+fun import_tac ctxt (thyname, thmname) =
+ if ! quick_and_dirty
+ then SkipProof.cheat_tac (ProofContext.theory_of ctxt)
+ else
+ fn th =>
+ let
+ val thy = ProofContext.theory_of ctxt
+ val prem = hd (prems_of th)
+ val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
+ val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
+ val int_thms = case ImportData.get thy of
+ NONE => fst (Replay.setup_int_thms thyname thy)
+ | SOME a => a
+ val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
+ val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
+ val thm = snd (ProofKernel.to_isa_thm hol4thm)
+ val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
+ val thm = equal_elim rew thm
+ val prew = ProofKernel.rewrite_hol4_term prem thy
+ val prem' = #2 (Logic.dest_equals (prop_of prew))
+ val _ = message ("Import proved " ^ Display.string_of_thm thm)
+ val thm = ProofKernel.disambiguate_frees thm
+ val _ = message ("Disambiguate: " ^ Display.string_of_thm thm)
+ in
+ case Shuffler.set_prop thy prem' [("",thm)] of
+ SOME (_,thm) =>
+ let
+ val _ = if prem' aconv (prop_of thm)
+ then message "import: Terms match up"
+ else message "import: Terms DO NOT match up"
+ val thm' = equal_elim (symmetric prew) thm
+ val res = bicompose true (false,thm',0) 1 th
+ in
+ res
+ end
+ | NONE => (message "import: set_prop didn't succeed"; no_tac th)
+ end
+
+val setup = Method.setup @{binding import}
+ (Scan.lift (Args.name -- Args.name) >>
+ (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
+ "import HOL4 theorem"
+
+end
+
--- a/src/HOL/Import/import_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-(* Title: HOL/Import/import_package.ML
- Author: Sebastian Skalberg (TU Muenchen)
-*)
-
-signature IMPORT_PACKAGE =
-sig
- val debug : bool ref
- val import_tac : Proof.context -> string * string -> tactic
- val setup : theory -> theory
-end
-
-structure ImportData = TheoryDataFun
-(
- type T = ProofKernel.thm option array option
- val empty = NONE
- val copy = I
- val extend = I
- fun merge _ _ = NONE
-)
-
-structure ImportPackage :> IMPORT_PACKAGE =
-struct
-
-val debug = ref false
-fun message s = if !debug then writeln s else ()
-
-fun import_tac ctxt (thyname, thmname) =
- if ! quick_and_dirty
- then SkipProof.cheat_tac (ProofContext.theory_of ctxt)
- else
- fn th =>
- let
- val thy = ProofContext.theory_of ctxt
- val prem = hd (prems_of th)
- val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
- val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
- val int_thms = case ImportData.get thy of
- NONE => fst (Replay.setup_int_thms thyname thy)
- | SOME a => a
- val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
- val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
- val thm = snd (ProofKernel.to_isa_thm hol4thm)
- val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
- val thm = equal_elim rew thm
- val prew = ProofKernel.rewrite_hol4_term prem thy
- val prem' = #2 (Logic.dest_equals (prop_of prew))
- val _ = message ("Import proved " ^ Display.string_of_thm thm)
- val thm = ProofKernel.disambiguate_frees thm
- val _ = message ("Disambiguate: " ^ Display.string_of_thm thm)
- in
- case Shuffler.set_prop thy prem' [("",thm)] of
- SOME (_,thm) =>
- let
- val _ = if prem' aconv (prop_of thm)
- then message "import: Terms match up"
- else message "import: Terms DO NOT match up"
- val thm' = equal_elim (symmetric prew) thm
- val res = bicompose true (false,thm',0) 1 th
- in
- res
- end
- | NONE => (message "import: set_prop didn't succeed"; no_tac th)
- end
-
-val setup = Method.setup @{binding import}
- (Scan.lift (Args.name -- Args.name) >>
- (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
- "import HOL4 theorem"
-
-end
-
--- a/src/HOL/Import/proof_kernel.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Import/proof_kernel.ML Tue Jun 23 15:48:55 2009 +0100
@@ -2021,7 +2021,7 @@
snd (get_defname thyname name thy)) thy1 names
fun new_name name = fst (get_defname thyname name thy1)
val names' = map (fn name => (new_name name,name,false)) names
- val (thy',res) = SpecificationPackage.add_specification NONE
+ val (thy',res) = Choice_Specification.add_specification NONE
names'
(thy1,th)
val _ = ImportRecorder.add_specification names' th
@@ -2091,7 +2091,7 @@
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
val ((_, typedef_info), thy') =
- TypedefPackage.add_typedef false (SOME (Binding.name thmname))
+ Typedef.add_typedef false (SOME (Binding.name thmname))
(Binding.name tycname, tnames, tsyn) c NONE (rtac th2 1) thy
val _ = ImportRecorder.add_typedef (SOME thmname) typ c NONE th2
@@ -2179,7 +2179,7 @@
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
val ((_, typedef_info), thy') =
- TypedefPackage.add_typedef false NONE (Binding.name tycname,tnames,tsyn) c
+ Typedef.add_typedef false NONE (Binding.name tycname,tnames,tsyn) c
(SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
val _ = ImportRecorder.add_typedef NONE typ c (SOME(rep_name,abs_name)) th2
val fulltyname = Sign.intern_type thy' tycname
--- a/src/HOL/Import/replay.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Import/replay.ML Tue Jun 23 15:48:55 2009 +0100
@@ -329,7 +329,7 @@
and rp (ThmEntry (thyname', thmname', aborted, History history)) thy = rps history thy
| rp (DeltaEntry ds) thy = fold delta ds thy
and delta (Specification (names, th)) thy =
- fst (SpecificationPackage.add_specification NONE names (thy,th_of thy th))
+ fst (Choice_Specification.add_specification NONE names (thy,th_of thy th))
| delta (Hol_mapping (thyname, thmname, isaname)) thy =
add_hol4_mapping thyname thmname isaname thy
| delta (Hol_pending (thyname, thmname, th)) thy =
@@ -344,7 +344,7 @@
| delta (Hol_theorem (thyname, thmname, th)) thy =
add_hol4_theorem thyname thmname ([], th_of thy th) thy
| delta (Typedef (thmname, (t, vs, mx), c, repabs, th)) thy =
- snd (TypedefPackage.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
+ snd (Typedef.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
(Option.map (pairself Binding.name) repabs) (rtac (th_of thy th) 1) thy)
| delta (Hol_type_mapping (thyname, tycname, fulltyname)) thy =
add_hol4_type_mapping thyname tycname true fulltyname thy
--- a/src/HOL/Inductive.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Inductive.thy Tue Jun 23 15:48:55 2009 +0100
@@ -7,18 +7,18 @@
theory Inductive
imports Lattices Sum_Type
uses
- ("Tools/inductive_package.ML")
+ ("Tools/inductive.ML")
"Tools/dseq.ML"
("Tools/inductive_codegen.ML")
- ("Tools/datatype_package/datatype_aux.ML")
- ("Tools/datatype_package/datatype_prop.ML")
- ("Tools/datatype_package/datatype_rep_proofs.ML")
- ("Tools/datatype_package/datatype_abs_proofs.ML")
- ("Tools/datatype_package/datatype_case.ML")
- ("Tools/datatype_package/datatype_package.ML")
- ("Tools/old_primrec_package.ML")
- ("Tools/primrec_package.ML")
- ("Tools/datatype_package/datatype_codegen.ML")
+ ("Tools/Datatype/datatype_aux.ML")
+ ("Tools/Datatype/datatype_prop.ML")
+ ("Tools/Datatype/datatype_rep_proofs.ML")
+ ("Tools/Datatype/datatype_abs_proofs.ML")
+ ("Tools/Datatype/datatype_case.ML")
+ ("Tools/Datatype/datatype.ML")
+ ("Tools/old_primrec.ML")
+ ("Tools/primrec.ML")
+ ("Tools/Datatype/datatype_codegen.ML")
begin
subsection {* Least and greatest fixed points *}
@@ -320,8 +320,8 @@
val le_fun_def = @{thm le_fun_def} RS @{thm eq_reflection}
*}
-use "Tools/inductive_package.ML"
-setup InductivePackage.setup
+use "Tools/inductive.ML"
+setup Inductive.setup
theorems [mono] =
imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj
@@ -335,18 +335,18 @@
text {* Package setup. *}
-use "Tools/datatype_package/datatype_aux.ML"
-use "Tools/datatype_package/datatype_prop.ML"
-use "Tools/datatype_package/datatype_rep_proofs.ML"
-use "Tools/datatype_package/datatype_abs_proofs.ML"
-use "Tools/datatype_package/datatype_case.ML"
-use "Tools/datatype_package/datatype_package.ML"
-setup DatatypePackage.setup
+use "Tools/Datatype/datatype_aux.ML"
+use "Tools/Datatype/datatype_prop.ML"
+use "Tools/Datatype/datatype_rep_proofs.ML"
+use "Tools/Datatype/datatype_abs_proofs.ML"
+use "Tools/Datatype/datatype_case.ML"
+use "Tools/Datatype/datatype.ML"
+setup Datatype.setup
-use "Tools/old_primrec_package.ML"
-use "Tools/primrec_package.ML"
+use "Tools/old_primrec.ML"
+use "Tools/primrec.ML"
-use "Tools/datatype_package/datatype_codegen.ML"
+use "Tools/Datatype/datatype_codegen.ML"
setup DatatypeCodegen.setup
use "Tools/inductive_codegen.ML"
@@ -364,7 +364,7 @@
fun fun_tr ctxt [cs] =
let
val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
- val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr
+ val ft = DatatypeCase.case_tr true Datatype.datatype_of_constr
ctxt [x, cs]
in lambda x ft end
in [("_lam_pats_syntax", fun_tr)] end
--- a/src/HOL/IntDiv.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/IntDiv.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1064,6 +1064,16 @@
lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)"
by (rule dvd_mod_imp_dvd) (* TODO: remove *)
+lemma dvd_imp_le_int: "(i::int) ~= 0 ==> d dvd i ==> abs d <= abs i"
+apply(auto simp:abs_if)
+ apply(clarsimp simp:dvd_def mult_less_0_iff)
+ using mult_le_cancel_left_neg[of _ "-1::int"]
+ apply(clarsimp simp:dvd_def mult_less_0_iff)
+ apply(clarsimp simp:dvd_def mult_less_0_iff
+ minus_mult_right simp del: mult_minus_right)
+apply(clarsimp simp:dvd_def mult_less_0_iff)
+done
+
lemma zdvd_not_zless: "0 < m ==> m < n ==> \<not> n dvd (m::int)"
apply (auto elim!: dvdE)
apply (subgoal_tac "0 < n")
--- a/src/HOL/IsaMakefile Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/IsaMakefile Tue Jun 23 15:48:55 2009 +0100
@@ -34,6 +34,7 @@
HOL-Modelcheck \
HOL-NanoJava \
HOL-Nominal-Examples \
+ HOL-NewNumberTheory \
HOL-NumberTheory \
HOL-Prolog \
HOL-SET-Protocol \
@@ -90,12 +91,12 @@
$(SRC)/Tools/IsaPlanner/zipper.ML \
$(SRC)/Tools/atomize_elim.ML \
$(SRC)/Tools/auto_solve.ML \
- $(SRC)/Tools/code/code_haskell.ML \
- $(SRC)/Tools/code/code_ml.ML \
- $(SRC)/Tools/code/code_preproc.ML \
- $(SRC)/Tools/code/code_printer.ML \
- $(SRC)/Tools/code/code_target.ML \
- $(SRC)/Tools/code/code_thingol.ML \
+ $(SRC)/Tools/Code/code_haskell.ML \
+ $(SRC)/Tools/Code/code_ml.ML \
+ $(SRC)/Tools/Code/code_preproc.ML \
+ $(SRC)/Tools/Code/code_printer.ML \
+ $(SRC)/Tools/Code/code_target.ML \
+ $(SRC)/Tools/Code/code_thingol.ML \
$(SRC)/Tools/coherent.ML \
$(SRC)/Tools/eqsubst.ML \
$(SRC)/Tools/induct.ML \
@@ -141,54 +142,54 @@
Sum_Type.thy \
Tools/arith_data.ML \
Tools/cnf_funcs.ML \
- Tools/datatype_package/datatype_abs_proofs.ML \
- Tools/datatype_package/datatype_aux.ML \
- Tools/datatype_package/datatype_case.ML \
- Tools/datatype_package/datatype_codegen.ML \
- Tools/datatype_package/datatype_package.ML \
- Tools/datatype_package/datatype_prop.ML \
- Tools/datatype_package/datatype_realizer.ML \
- Tools/datatype_package/datatype_rep_proofs.ML \
+ Tools/Datatype/datatype_abs_proofs.ML \
+ Tools/Datatype/datatype_aux.ML \
+ Tools/Datatype/datatype_case.ML \
+ Tools/Datatype/datatype_codegen.ML \
+ Tools/Datatype/datatype.ML \
+ Tools/Datatype/datatype_prop.ML \
+ Tools/Datatype/datatype_realizer.ML \
+ Tools/Datatype/datatype_rep_proofs.ML \
Tools/dseq.ML \
- Tools/function_package/auto_term.ML \
- Tools/function_package/context_tree.ML \
- Tools/function_package/decompose.ML \
- Tools/function_package/descent.ML \
- Tools/function_package/fundef_common.ML \
- Tools/function_package/fundef_core.ML \
- Tools/function_package/fundef_datatype.ML \
- Tools/function_package/fundef_lib.ML \
- Tools/function_package/fundef_package.ML \
- Tools/function_package/induction_scheme.ML \
- Tools/function_package/inductive_wrap.ML \
- Tools/function_package/lexicographic_order.ML \
- Tools/function_package/measure_functions.ML \
- Tools/function_package/mutual.ML \
- Tools/function_package/pattern_split.ML \
- Tools/function_package/scnp_reconstruct.ML \
- Tools/function_package/scnp_solve.ML \
- Tools/function_package/size.ML \
- Tools/function_package/sum_tree.ML \
- Tools/function_package/termination.ML \
+ Tools/Function/auto_term.ML \
+ Tools/Function/context_tree.ML \
+ Tools/Function/decompose.ML \
+ Tools/Function/descent.ML \
+ Tools/Function/fundef_common.ML \
+ Tools/Function/fundef_core.ML \
+ Tools/Function/fundef_datatype.ML \
+ Tools/Function/fundef_lib.ML \
+ Tools/Function/fundef.ML \
+ Tools/Function/induction_scheme.ML \
+ Tools/Function/inductive_wrap.ML \
+ Tools/Function/lexicographic_order.ML \
+ Tools/Function/measure_functions.ML \
+ Tools/Function/mutual.ML \
+ Tools/Function/pattern_split.ML \
+ Tools/Function/scnp_reconstruct.ML \
+ Tools/Function/scnp_solve.ML \
+ Tools/Function/size.ML \
+ Tools/Function/sum_tree.ML \
+ Tools/Function/termination.ML \
Tools/inductive_codegen.ML \
- Tools/inductive_package.ML \
+ Tools/inductive.ML \
Tools/inductive_realizer.ML \
- Tools/inductive_set_package.ML \
+ Tools/inductive_set.ML \
Tools/lin_arith.ML \
Tools/nat_arith.ML \
- Tools/old_primrec_package.ML \
- Tools/primrec_package.ML \
+ Tools/old_primrec.ML \
+ Tools/primrec.ML \
Tools/prop_logic.ML \
- Tools/record_package.ML \
+ Tools/record.ML \
Tools/refute.ML \
Tools/refute_isar.ML \
Tools/rewrite_hol_proof.ML \
Tools/sat_funcs.ML \
Tools/sat_solver.ML \
Tools/split_rule.ML \
- Tools/typecopy_package.ML \
+ Tools/typecopy.ML \
Tools/typedef_codegen.ML \
- Tools/typedef_package.ML \
+ Tools/typedef.ML \
Transitive_Closure.thy \
Typedef.thy \
Wellfounded.thy \
@@ -250,13 +251,13 @@
Tools/Qelim/generated_cooper.ML \
Tools/Qelim/presburger.ML \
Tools/Qelim/qelim.ML \
- Tools/recdef_package.ML \
+ Tools/recdef.ML \
Tools/res_atp.ML \
Tools/res_axioms.ML \
Tools/res_clause.ML \
Tools/res_hol_clause.ML \
Tools/res_reconstruct.ML \
- Tools/specification_package.ML \
+ Tools/choice_specification.ML \
Tools/string_code.ML \
Tools/string_syntax.ML \
Tools/TFL/casesplit.ML \
@@ -284,6 +285,7 @@
Ln.thy \
Log.thy \
MacLaurin.thy \
+ NatTransfer.thy \
NthRoot.thy \
SEQ.thy \
Series.thy \
@@ -300,6 +302,7 @@
Real.thy \
RealVector.thy \
Tools/float_syntax.ML \
+ Tools/transfer_data.ML \
Tools/Qelim/ferrante_rackoff_data.ML \
Tools/Qelim/ferrante_rackoff.ML \
Tools/Qelim/langford_data.ML \
@@ -322,9 +325,10 @@
Library/FuncSet.thy Library/Permutations.thy Library/Determinants.thy\
Library/Bit.thy Library/Topology_Euclidean_Space.thy \
Library/Finite_Cartesian_Product.thy \
- Library/FrechetDeriv.thy \
+ Library/FrechetDeriv.thy Library/Fraction_Field.thy\
Library/Fundamental_Theorem_Algebra.thy \
Library/Inner_Product.thy Library/Lattice_Syntax.thy \
+ Library/Legacy_GCD.thy \
Library/Library.thy Library/List_Prefix.thy Library/State_Monad.thy \
Library/Nat_Int_Bij.thy Library/Multiset.thy Library/Permutation.thy \
Library/Primes.thy Library/Pocklington.thy Library/Quotient.thy \
@@ -420,12 +424,12 @@
IMPORTER_FILES = Import/lazy_seq.ML Import/proof_kernel.ML Import/replay.ML \
Import/shuffler.ML Import/MakeEqual.thy Import/HOL4Setup.thy \
Import/HOL4Syntax.thy Import/HOL4Compat.thy Import/import_syntax.ML \
- Import/hol4rews.ML Import/import_package.ML Import/ROOT.ML
+ Import/hol4rews.ML Import/import.ML Import/ROOT.ML
IMPORTER_HOLLIGHT_FILES = Import/proof_kernel.ML Import/replay.ML \
Import/shuffler.ML Import/MakeEqual.thy Import/HOL4Setup.thy \
Import/HOL4Syntax.thy Import/HOLLightCompat.thy Import/import_syntax.ML \
- Import/hol4rews.ML Import/import_package.ML Import/ROOT.ML
+ Import/hol4rews.ML Import/import.ML Import/ROOT.ML
HOL-Import: HOL $(LOG)/HOL-Import.gz
@@ -486,6 +490,22 @@
@cd Import/HOLLight; $(ISABELLE_TOOL) usedir -b $(OUT)/HOL HOLLight
+## HOL-NewNumberTheory
+
+HOL-NewNumberTheory: HOL $(LOG)/HOL-NewNumberTheory.gz
+
+$(LOG)/HOL-NewNumberTheory.gz: $(OUT)/HOL $(ALGEBRA_DEPENDENCIES) \
+ Library/Multiset.thy \
+ NewNumberTheory/Binomial.thy \
+ NewNumberTheory/Cong.thy \
+ NewNumberTheory/Fib.thy \
+ NewNumberTheory/MiscAlgebra.thy \
+ NewNumberTheory/Residues.thy \
+ NewNumberTheory/UniqueFactorization.thy \
+ NewNumberTheory/ROOT.ML
+ @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL NewNumberTheory
+
+
## HOL-NumberTheory
HOL-NumberTheory: HOL $(LOG)/HOL-NumberTheory.gz
@@ -550,22 +570,46 @@
HOL-Algebra: HOL $(LOG)/HOL-Algebra.gz
-$(LOG)/HOL-Algebra.gz: $(OUT)/HOL Algebra/ROOT.ML \
- Library/Binomial.thy Library/FuncSet.thy \
- Library/Multiset.thy Library/Permutation.thy Library/Primes.thy \
- Algebra/AbelCoset.thy Algebra/Bij.thy Algebra/Congruence.thy \
- Algebra/Coset.thy Algebra/Divisibility.thy Algebra/Exponent.thy \
- Algebra/FiniteProduct.thy \
- Algebra/Group.thy Algebra/Ideal.thy Algebra/IntRing.thy \
- Algebra/Lattice.thy Algebra/Module.thy Algebra/QuotRing.thy \
- Algebra/Ring.thy Algebra/RingHom.thy Algebra/Sylow.thy \
- Algebra/UnivPoly.thy Algebra/abstract/Abstract.thy \
- Algebra/abstract/Factor.thy Algebra/abstract/Field.thy \
- Algebra/abstract/Ideal2.thy Algebra/abstract/PID.thy \
- Algebra/abstract/Ring2.thy Algebra/abstract/RingHomo.thy \
- Algebra/document/root.tex Algebra/poly/LongDiv.thy \
- Algebra/poly/PolyHomo.thy Algebra/poly/Polynomial.thy \
- Algebra/poly/UnivPoly2.thy Algebra/ringsimp.ML
+ALGEBRA_DEPENDENCIES = $(OUT)/HOL \
+ Algebra/ROOT.ML \
+ Library/Binomial.thy \
+ Library/FuncSet.thy \
+ Library/Multiset.thy \
+ Library/Permutation.thy \
+ Library/Primes.thy \
+ Algebra/AbelCoset.thy \
+ Algebra/Bij.thy \
+ Algebra/Congruence.thy \
+ Algebra/Coset.thy \
+ Algebra/Divisibility.thy \
+ Algebra/Exponent.thy \
+ Algebra/FiniteProduct.thy \
+ Algebra/Group.thy \
+ Algebra/Ideal.thy \
+ Algebra/IntRing.thy \
+ Algebra/Lattice.thy \
+ Algebra/Module.thy \
+ Algebra/QuotRing.thy \
+ Algebra/Ring.thy \
+ Algebra/RingHom.thy \
+ Algebra/Sylow.thy \
+ Algebra/UnivPoly.thy \
+ Algebra/abstract/Abstract.thy \
+ Algebra/abstract/Factor.thy \
+ Algebra/abstract/Field.thy \
+ Algebra/abstract/Ideal2.thy \
+ Algebra/abstract/PID.thy \
+ Algebra/abstract/Ring2.thy \
+ Algebra/abstract/RingHomo.thy \
+ Algebra/document/root.tex \
+ Algebra/document/root.tex \
+ Algebra/poly/LongDiv.thy \
+ Algebra/poly/PolyHomo.thy \
+ Algebra/poly/Polynomial.thy \
+ Algebra/poly/UnivPoly2.thy \
+ Algebra/ringsimp.ML
+
+$(LOG)/HOL-Algebra.gz: $(ALGEBRA_DEPENDENCIES)
@cd Algebra; $(ISABELLE_TOOL) usedir -b -g true -V outline=/proof,/ML $(OUT)/HOL HOL-Algebra
@@ -869,13 +913,13 @@
HOL-Isar_examples: HOL $(LOG)/HOL-Isar_examples.gz
-$(LOG)/HOL-Isar_examples.gz: $(OUT)/HOL Isar_examples/BasicLogic.thy \
+$(LOG)/HOL-Isar_examples.gz: $(OUT)/HOL Isar_examples/Basic_Logic.thy \
Isar_examples/Cantor.thy Isar_examples/Drinker.thy \
- Isar_examples/ExprCompiler.thy Isar_examples/Fibonacci.thy \
+ Isar_examples/Expr_Compiler.thy Isar_examples/Fibonacci.thy \
Isar_examples/Group.thy Isar_examples/Hoare.thy \
- Isar_examples/HoareEx.thy Isar_examples/KnasterTarski.thy \
- Isar_examples/MutilatedCheckerboard.thy \
- Isar_examples/NestedDatatype.thy Isar_examples/Peirce.thy \
+ Isar_examples/Hoare_Ex.thy Isar_examples/Knaster_Tarski.thy \
+ Isar_examples/Mutilated_Checkerboard.thy \
+ Isar_examples/Nested_Datatype.thy Isar_examples/Peirce.thy \
Isar_examples/Puzzle.thy Isar_examples/Summation.thy \
Isar_examples/ROOT.ML Isar_examples/document/proof.sty \
Isar_examples/document/root.bib Isar_examples/document/root.tex \
@@ -965,7 +1009,7 @@
Nominal/nominal_induct.ML \
Nominal/nominal_inductive.ML \
Nominal/nominal_inductive2.ML \
- Nominal/nominal_package.ML \
+ Nominal/nominal.ML \
Nominal/nominal_permeq.ML \
Nominal/nominal_primrec.ML \
Nominal/nominal_thmdecls.ML \
--- a/src/HOL/Isar_examples/BasicLogic.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,447 +0,0 @@
-(* Title: HOL/Isar_examples/BasicLogic.thy
- ID: $Id$
- Author: Markus Wenzel, TU Muenchen
-
-Basic propositional and quantifier reasoning.
-*)
-
-header {* Basic logical reasoning *}
-
-theory BasicLogic imports Main begin
-
-
-subsection {* Pure backward reasoning *}
-
-text {*
- In order to get a first idea of how Isabelle/Isar proof documents
- may look like, we consider the propositions @{text I}, @{text K},
- and @{text S}. The following (rather explicit) proofs should
- require little extra explanations.
-*}
-
-lemma I: "A --> A"
-proof
- assume A
- show A by fact
-qed
-
-lemma K: "A --> B --> A"
-proof
- assume A
- show "B --> A"
- proof
- show A by fact
- qed
-qed
-
-lemma S: "(A --> B --> C) --> (A --> B) --> A --> C"
-proof
- assume "A --> B --> C"
- show "(A --> B) --> A --> C"
- proof
- assume "A --> B"
- show "A --> C"
- proof
- assume A
- show C
- proof (rule mp)
- show "B --> C" by (rule mp) fact+
- show B by (rule mp) fact+
- qed
- qed
- qed
-qed
-
-text {*
- Isar provides several ways to fine-tune the reasoning, avoiding
- excessive detail. Several abbreviated language elements are
- available, enabling the writer to express proofs in a more concise
- way, even without referring to any automated proof tools yet.
-
- First of all, proof by assumption may be abbreviated as a single
- dot.
-*}
-
-lemma "A --> A"
-proof
- assume A
- show A by fact+
-qed
-
-text {*
- In fact, concluding any (sub-)proof already involves solving any
- remaining goals by assumption\footnote{This is not a completely
- trivial operation, as proof by assumption may involve full
- higher-order unification.}. Thus we may skip the rather vacuous
- body of the above proof as well.
-*}
-
-lemma "A --> A"
-proof
-qed
-
-text {*
- Note that the \isacommand{proof} command refers to the @{text rule}
- method (without arguments) by default. Thus it implicitly applies a
- single rule, as determined from the syntactic form of the statements
- involved. The \isacommand{by} command abbreviates any proof with
- empty body, so the proof may be further pruned.
-*}
-
-lemma "A --> A"
- by rule
-
-text {*
- Proof by a single rule may be abbreviated as double-dot.
-*}
-
-lemma "A --> A" ..
-
-text {*
- Thus we have arrived at an adequate representation of the proof of a
- tautology that holds by a single standard rule.\footnote{Apparently,
- the rule here is implication introduction.}
-*}
-
-text {*
- Let us also reconsider @{text K}. Its statement is composed of
- iterated connectives. Basic decomposition is by a single rule at a
- time, which is why our first version above was by nesting two
- proofs.
-
- The @{text intro} proof method repeatedly decomposes a goal's
- conclusion.\footnote{The dual method is @{text elim}, acting on a
- goal's premises.}
-*}
-
-lemma "A --> B --> A"
-proof (intro impI)
- assume A
- show A by fact
-qed
-
-text {*
- Again, the body may be collapsed.
-*}
-
-lemma "A --> B --> A"
- by (intro impI)
-
-text {*
- Just like @{text rule}, the @{text intro} and @{text elim} proof
- methods pick standard structural rules, in case no explicit
- arguments are given. While implicit rules are usually just fine for
- single rule application, this may go too far with iteration. Thus
- in practice, @{text intro} and @{text elim} would be typically
- restricted to certain structures by giving a few rules only, e.g.\
- \isacommand{proof}~@{text "(intro impI allI)"} to strip implications
- and universal quantifiers.
-
- Such well-tuned iterated decomposition of certain structures is the
- prime application of @{text intro} and @{text elim}. In contrast,
- terminal steps that solve a goal completely are usually performed by
- actual automated proof methods (such as \isacommand{by}~@{text
- blast}.
-*}
-
-
-subsection {* Variations of backward vs.\ forward reasoning *}
-
-text {*
- Certainly, any proof may be performed in backward-style only. On
- the other hand, small steps of reasoning are often more naturally
- expressed in forward-style. Isar supports both backward and forward
- reasoning as a first-class concept. In order to demonstrate the
- difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}.
-
- The first version is purely backward.
-*}
-
-lemma "A & B --> B & A"
-proof
- assume "A & B"
- show "B & A"
- proof
- show B by (rule conjunct2) fact
- show A by (rule conjunct1) fact
- qed
-qed
-
-text {*
- Above, the @{text "conjunct_1/2"} projection rules had to be named
- explicitly, since the goals @{text B} and @{text A} did not provide
- any structural clue. This may be avoided using \isacommand{from} to
- focus on the @{text "A \<and> B"} assumption as the current facts,
- enabling the use of double-dot proofs. Note that \isacommand{from}
- already does forward-chaining, involving the \name{conjE} rule here.
-*}
-
-lemma "A & B --> B & A"
-proof
- assume "A & B"
- show "B & A"
- proof
- from `A & B` show B ..
- from `A & B` show A ..
- qed
-qed
-
-text {*
- In the next version, we move the forward step one level upwards.
- Forward-chaining from the most recent facts is indicated by the
- \isacommand{then} command. Thus the proof of @{text "B \<and> A"} from
- @{text "A \<and> B"} actually becomes an elimination, rather than an
- introduction. The resulting proof structure directly corresponds to
- that of the @{text conjE} rule, including the repeated goal
- proposition that is abbreviated as @{text ?thesis} below.
-*}
-
-lemma "A & B --> B & A"
-proof
- assume "A & B"
- then show "B & A"
- proof -- {* rule @{text conjE} of @{text "A \<and> B"} *}
- assume B A
- then show ?thesis .. -- {* rule @{text conjI} of @{text "B \<and> A"} *}
- qed
-qed
-
-text {*
- In the subsequent version we flatten the structure of the main body
- by doing forward reasoning all the time. Only the outermost
- decomposition step is left as backward.
-*}
-
-lemma "A & B --> B & A"
-proof
- assume "A & B"
- from `A & B` have A ..
- from `A & B` have B ..
- from `B` `A` show "B & A" ..
-qed
-
-text {*
- We can still push forward-reasoning a bit further, even at the risk
- of getting ridiculous. Note that we force the initial proof step to
- do nothing here, by referring to the ``-'' proof method.
-*}
-
-lemma "A & B --> B & A"
-proof -
- {
- assume "A & B"
- from `A & B` have A ..
- from `A & B` have B ..
- from `B` `A` have "B & A" ..
- }
- then show ?thesis .. -- {* rule \name{impI} *}
-qed
-
-text {*
- \medskip With these examples we have shifted through a whole range
- from purely backward to purely forward reasoning. Apparently, in
- the extreme ends we get slightly ill-structured proofs, which also
- require much explicit naming of either rules (backward) or local
- facts (forward).
-
- The general lesson learned here is that good proof style would
- achieve just the \emph{right} balance of top-down backward
- decomposition, and bottom-up forward composition. In general, there
- is no single best way to arrange some pieces of formal reasoning, of
- course. Depending on the actual applications, the intended audience
- etc., rules (and methods) on the one hand vs.\ facts on the other
- hand have to be emphasized in an appropriate way. This requires the
- proof writer to develop good taste, and some practice, of course.
-*}
-
-text {*
- For our example the most appropriate way of reasoning is probably
- the middle one, with conjunction introduction done after
- elimination.
-*}
-
-lemma "A & B --> B & A"
-proof
- assume "A & B"
- then show "B & A"
- proof
- assume B A
- then show ?thesis ..
- qed
-qed
-
-
-
-subsection {* A few examples from ``Introduction to Isabelle'' *}
-
-text {*
- We rephrase some of the basic reasoning examples of
- \cite{isabelle-intro}, using HOL rather than FOL.
-*}
-
-subsubsection {* A propositional proof *}
-
-text {*
- We consider the proposition @{text "P \<or> P \<longrightarrow> P"}. The proof below
- involves forward-chaining from @{text "P \<or> P"}, followed by an
- explicit case-analysis on the two \emph{identical} cases.
-*}
-
-lemma "P | P --> P"
-proof
- assume "P | P"
- then show P
- proof -- {*
- rule @{text disjE}: \smash{$\infer{C}{A \disj B & \infer*{C}{[A]} & \infer*{C}{[B]}}$}
- *}
- assume P show P by fact
- next
- assume P show P by fact
- qed
-qed
-
-text {*
- Case splits are \emph{not} hardwired into the Isar language as a
- special feature. The \isacommand{next} command used to separate the
- cases above is just a short form of managing block structure.
-
- \medskip In general, applying proof methods may split up a goal into
- separate ``cases'', i.e.\ new subgoals with individual local
- assumptions. The corresponding proof text typically mimics this by
- establishing results in appropriate contexts, separated by blocks.
-
- In order to avoid too much explicit parentheses, the Isar system
- implicitly opens an additional block for any new goal, the
- \isacommand{next} statement then closes one block level, opening a
- new one. The resulting behavior is what one would expect from
- separating cases, only that it is more flexible. E.g.\ an induction
- base case (which does not introduce local assumptions) would
- \emph{not} require \isacommand{next} to separate the subsequent step
- case.
-
- \medskip In our example the situation is even simpler, since the two
- cases actually coincide. Consequently the proof may be rephrased as
- follows.
-*}
-
-lemma "P | P --> P"
-proof
- assume "P | P"
- then show P
- proof
- assume P
- show P by fact
- show P by fact
- qed
-qed
-
-text {*
- Again, the rather vacuous body of the proof may be collapsed. Thus
- the case analysis degenerates into two assumption steps, which are
- implicitly performed when concluding the single rule step of the
- double-dot proof as follows.
-*}
-
-lemma "P | P --> P"
-proof
- assume "P | P"
- then show P ..
-qed
-
-
-subsubsection {* A quantifier proof *}
-
-text {*
- To illustrate quantifier reasoning, let us prove @{text "(\<exists>x. P (f
- x)) \<longrightarrow> (\<exists>y. P y)"}. Informally, this holds because any @{text a}
- with @{text "P (f a)"} may be taken as a witness for the second
- existential statement.
-
- The first proof is rather verbose, exhibiting quite a lot of
- (redundant) detail. It gives explicit rules, even with some
- instantiation. Furthermore, we encounter two new language elements:
- the \isacommand{fix} command augments the context by some new
- ``arbitrary, but fixed'' element; the \isacommand{is} annotation
- binds term abbreviations by higher-order pattern matching.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
- assume "EX x. P (f x)"
- then show "EX y. P y"
- proof (rule exE) -- {*
- rule \name{exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$}
- *}
- fix a
- assume "P (f a)" (is "P ?witness")
- then show ?thesis by (rule exI [of P ?witness])
- qed
-qed
-
-text {*
- While explicit rule instantiation may occasionally improve
- readability of certain aspects of reasoning, it is usually quite
- redundant. Above, the basic proof outline gives already enough
- structural clues for the system to infer both the rules and their
- instances (by higher-order unification). Thus we may as well prune
- the text as follows.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
- assume "EX x. P (f x)"
- then show "EX y. P y"
- proof
- fix a
- assume "P (f a)"
- then show ?thesis ..
- qed
-qed
-
-text {*
- Explicit @{text \<exists>}-elimination as seen above can become quite
- cumbersome in practice. The derived Isar language element
- ``\isakeyword{obtain}'' provides a more handsome way to do
- generalized existence reasoning.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
- assume "EX x. P (f x)"
- then obtain a where "P (f a)" ..
- then show "EX y. P y" ..
-qed
-
-text {*
- Technically, \isakeyword{obtain} is similar to \isakeyword{fix} and
- \isakeyword{assume} together with a soundness proof of the
- elimination involved. Thus it behaves similar to any other forward
- proof element. Also note that due to the nature of general
- existence reasoning involved here, any result exported from the
- context of an \isakeyword{obtain} statement may \emph{not} refer to
- the parameters introduced there.
-*}
-
-
-
-subsubsection {* Deriving rules in Isabelle *}
-
-text {*
- We derive the conjunction elimination rule from the corresponding
- projections. The proof is quite straight-forward, since
- Isabelle/Isar supports non-atomic goals and assumptions fully
- transparently.
-*}
-
-theorem conjE: "A & B ==> (A ==> B ==> C) ==> C"
-proof -
- assume "A & B"
- assume r: "A ==> B ==> C"
- show C
- proof (rule r)
- show A by (rule conjunct1) fact
- show B by (rule conjunct2) fact
- qed
-qed
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Basic_Logic.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,448 @@
+(* Title: HOL/Isar_examples/Basic_Logic.thy
+ Author: Markus Wenzel, TU Muenchen
+
+Basic propositional and quantifier reasoning.
+*)
+
+header {* Basic logical reasoning *}
+
+theory Basic_Logic
+imports Main
+begin
+
+
+subsection {* Pure backward reasoning *}
+
+text {*
+ In order to get a first idea of how Isabelle/Isar proof documents
+ may look like, we consider the propositions @{text I}, @{text K},
+ and @{text S}. The following (rather explicit) proofs should
+ require little extra explanations.
+*}
+
+lemma I: "A --> A"
+proof
+ assume A
+ show A by fact
+qed
+
+lemma K: "A --> B --> A"
+proof
+ assume A
+ show "B --> A"
+ proof
+ show A by fact
+ qed
+qed
+
+lemma S: "(A --> B --> C) --> (A --> B) --> A --> C"
+proof
+ assume "A --> B --> C"
+ show "(A --> B) --> A --> C"
+ proof
+ assume "A --> B"
+ show "A --> C"
+ proof
+ assume A
+ show C
+ proof (rule mp)
+ show "B --> C" by (rule mp) fact+
+ show B by (rule mp) fact+
+ qed
+ qed
+ qed
+qed
+
+text {*
+ Isar provides several ways to fine-tune the reasoning, avoiding
+ excessive detail. Several abbreviated language elements are
+ available, enabling the writer to express proofs in a more concise
+ way, even without referring to any automated proof tools yet.
+
+ First of all, proof by assumption may be abbreviated as a single
+ dot.
+*}
+
+lemma "A --> A"
+proof
+ assume A
+ show A by fact+
+qed
+
+text {*
+ In fact, concluding any (sub-)proof already involves solving any
+ remaining goals by assumption\footnote{This is not a completely
+ trivial operation, as proof by assumption may involve full
+ higher-order unification.}. Thus we may skip the rather vacuous
+ body of the above proof as well.
+*}
+
+lemma "A --> A"
+proof
+qed
+
+text {*
+ Note that the \isacommand{proof} command refers to the @{text rule}
+ method (without arguments) by default. Thus it implicitly applies a
+ single rule, as determined from the syntactic form of the statements
+ involved. The \isacommand{by} command abbreviates any proof with
+ empty body, so the proof may be further pruned.
+*}
+
+lemma "A --> A"
+ by rule
+
+text {*
+ Proof by a single rule may be abbreviated as double-dot.
+*}
+
+lemma "A --> A" ..
+
+text {*
+ Thus we have arrived at an adequate representation of the proof of a
+ tautology that holds by a single standard rule.\footnote{Apparently,
+ the rule here is implication introduction.}
+*}
+
+text {*
+ Let us also reconsider @{text K}. Its statement is composed of
+ iterated connectives. Basic decomposition is by a single rule at a
+ time, which is why our first version above was by nesting two
+ proofs.
+
+ The @{text intro} proof method repeatedly decomposes a goal's
+ conclusion.\footnote{The dual method is @{text elim}, acting on a
+ goal's premises.}
+*}
+
+lemma "A --> B --> A"
+proof (intro impI)
+ assume A
+ show A by fact
+qed
+
+text {*
+ Again, the body may be collapsed.
+*}
+
+lemma "A --> B --> A"
+ by (intro impI)
+
+text {*
+ Just like @{text rule}, the @{text intro} and @{text elim} proof
+ methods pick standard structural rules, in case no explicit
+ arguments are given. While implicit rules are usually just fine for
+ single rule application, this may go too far with iteration. Thus
+ in practice, @{text intro} and @{text elim} would be typically
+ restricted to certain structures by giving a few rules only, e.g.\
+ \isacommand{proof}~@{text "(intro impI allI)"} to strip implications
+ and universal quantifiers.
+
+ Such well-tuned iterated decomposition of certain structures is the
+ prime application of @{text intro} and @{text elim}. In contrast,
+ terminal steps that solve a goal completely are usually performed by
+ actual automated proof methods (such as \isacommand{by}~@{text
+ blast}.
+*}
+
+
+subsection {* Variations of backward vs.\ forward reasoning *}
+
+text {*
+ Certainly, any proof may be performed in backward-style only. On
+ the other hand, small steps of reasoning are often more naturally
+ expressed in forward-style. Isar supports both backward and forward
+ reasoning as a first-class concept. In order to demonstrate the
+ difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}.
+
+ The first version is purely backward.
+*}
+
+lemma "A & B --> B & A"
+proof
+ assume "A & B"
+ show "B & A"
+ proof
+ show B by (rule conjunct2) fact
+ show A by (rule conjunct1) fact
+ qed
+qed
+
+text {*
+ Above, the @{text "conjunct_1/2"} projection rules had to be named
+ explicitly, since the goals @{text B} and @{text A} did not provide
+ any structural clue. This may be avoided using \isacommand{from} to
+ focus on the @{text "A \<and> B"} assumption as the current facts,
+ enabling the use of double-dot proofs. Note that \isacommand{from}
+ already does forward-chaining, involving the \name{conjE} rule here.
+*}
+
+lemma "A & B --> B & A"
+proof
+ assume "A & B"
+ show "B & A"
+ proof
+ from `A & B` show B ..
+ from `A & B` show A ..
+ qed
+qed
+
+text {*
+ In the next version, we move the forward step one level upwards.
+ Forward-chaining from the most recent facts is indicated by the
+ \isacommand{then} command. Thus the proof of @{text "B \<and> A"} from
+ @{text "A \<and> B"} actually becomes an elimination, rather than an
+ introduction. The resulting proof structure directly corresponds to
+ that of the @{text conjE} rule, including the repeated goal
+ proposition that is abbreviated as @{text ?thesis} below.
+*}
+
+lemma "A & B --> B & A"
+proof
+ assume "A & B"
+ then show "B & A"
+ proof -- {* rule @{text conjE} of @{text "A \<and> B"} *}
+ assume B A
+ then show ?thesis .. -- {* rule @{text conjI} of @{text "B \<and> A"} *}
+ qed
+qed
+
+text {*
+ In the subsequent version we flatten the structure of the main body
+ by doing forward reasoning all the time. Only the outermost
+ decomposition step is left as backward.
+*}
+
+lemma "A & B --> B & A"
+proof
+ assume "A & B"
+ from `A & B` have A ..
+ from `A & B` have B ..
+ from `B` `A` show "B & A" ..
+qed
+
+text {*
+ We can still push forward-reasoning a bit further, even at the risk
+ of getting ridiculous. Note that we force the initial proof step to
+ do nothing here, by referring to the ``-'' proof method.
+*}
+
+lemma "A & B --> B & A"
+proof -
+ {
+ assume "A & B"
+ from `A & B` have A ..
+ from `A & B` have B ..
+ from `B` `A` have "B & A" ..
+ }
+ then show ?thesis .. -- {* rule \name{impI} *}
+qed
+
+text {*
+ \medskip With these examples we have shifted through a whole range
+ from purely backward to purely forward reasoning. Apparently, in
+ the extreme ends we get slightly ill-structured proofs, which also
+ require much explicit naming of either rules (backward) or local
+ facts (forward).
+
+ The general lesson learned here is that good proof style would
+ achieve just the \emph{right} balance of top-down backward
+ decomposition, and bottom-up forward composition. In general, there
+ is no single best way to arrange some pieces of formal reasoning, of
+ course. Depending on the actual applications, the intended audience
+ etc., rules (and methods) on the one hand vs.\ facts on the other
+ hand have to be emphasized in an appropriate way. This requires the
+ proof writer to develop good taste, and some practice, of course.
+*}
+
+text {*
+ For our example the most appropriate way of reasoning is probably
+ the middle one, with conjunction introduction done after
+ elimination.
+*}
+
+lemma "A & B --> B & A"
+proof
+ assume "A & B"
+ then show "B & A"
+ proof
+ assume B A
+ then show ?thesis ..
+ qed
+qed
+
+
+
+subsection {* A few examples from ``Introduction to Isabelle'' *}
+
+text {*
+ We rephrase some of the basic reasoning examples of
+ \cite{isabelle-intro}, using HOL rather than FOL.
+*}
+
+subsubsection {* A propositional proof *}
+
+text {*
+ We consider the proposition @{text "P \<or> P \<longrightarrow> P"}. The proof below
+ involves forward-chaining from @{text "P \<or> P"}, followed by an
+ explicit case-analysis on the two \emph{identical} cases.
+*}
+
+lemma "P | P --> P"
+proof
+ assume "P | P"
+ then show P
+ proof -- {*
+ rule @{text disjE}: \smash{$\infer{C}{A \disj B & \infer*{C}{[A]} & \infer*{C}{[B]}}$}
+ *}
+ assume P show P by fact
+ next
+ assume P show P by fact
+ qed
+qed
+
+text {*
+ Case splits are \emph{not} hardwired into the Isar language as a
+ special feature. The \isacommand{next} command used to separate the
+ cases above is just a short form of managing block structure.
+
+ \medskip In general, applying proof methods may split up a goal into
+ separate ``cases'', i.e.\ new subgoals with individual local
+ assumptions. The corresponding proof text typically mimics this by
+ establishing results in appropriate contexts, separated by blocks.
+
+ In order to avoid too much explicit parentheses, the Isar system
+ implicitly opens an additional block for any new goal, the
+ \isacommand{next} statement then closes one block level, opening a
+ new one. The resulting behavior is what one would expect from
+ separating cases, only that it is more flexible. E.g.\ an induction
+ base case (which does not introduce local assumptions) would
+ \emph{not} require \isacommand{next} to separate the subsequent step
+ case.
+
+ \medskip In our example the situation is even simpler, since the two
+ cases actually coincide. Consequently the proof may be rephrased as
+ follows.
+*}
+
+lemma "P | P --> P"
+proof
+ assume "P | P"
+ then show P
+ proof
+ assume P
+ show P by fact
+ show P by fact
+ qed
+qed
+
+text {*
+ Again, the rather vacuous body of the proof may be collapsed. Thus
+ the case analysis degenerates into two assumption steps, which are
+ implicitly performed when concluding the single rule step of the
+ double-dot proof as follows.
+*}
+
+lemma "P | P --> P"
+proof
+ assume "P | P"
+ then show P ..
+qed
+
+
+subsubsection {* A quantifier proof *}
+
+text {*
+ To illustrate quantifier reasoning, let us prove @{text "(\<exists>x. P (f
+ x)) \<longrightarrow> (\<exists>y. P y)"}. Informally, this holds because any @{text a}
+ with @{text "P (f a)"} may be taken as a witness for the second
+ existential statement.
+
+ The first proof is rather verbose, exhibiting quite a lot of
+ (redundant) detail. It gives explicit rules, even with some
+ instantiation. Furthermore, we encounter two new language elements:
+ the \isacommand{fix} command augments the context by some new
+ ``arbitrary, but fixed'' element; the \isacommand{is} annotation
+ binds term abbreviations by higher-order pattern matching.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+ assume "EX x. P (f x)"
+ then show "EX y. P y"
+ proof (rule exE) -- {*
+ rule \name{exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$}
+ *}
+ fix a
+ assume "P (f a)" (is "P ?witness")
+ then show ?thesis by (rule exI [of P ?witness])
+ qed
+qed
+
+text {*
+ While explicit rule instantiation may occasionally improve
+ readability of certain aspects of reasoning, it is usually quite
+ redundant. Above, the basic proof outline gives already enough
+ structural clues for the system to infer both the rules and their
+ instances (by higher-order unification). Thus we may as well prune
+ the text as follows.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+ assume "EX x. P (f x)"
+ then show "EX y. P y"
+ proof
+ fix a
+ assume "P (f a)"
+ then show ?thesis ..
+ qed
+qed
+
+text {*
+ Explicit @{text \<exists>}-elimination as seen above can become quite
+ cumbersome in practice. The derived Isar language element
+ ``\isakeyword{obtain}'' provides a more handsome way to do
+ generalized existence reasoning.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+ assume "EX x. P (f x)"
+ then obtain a where "P (f a)" ..
+ then show "EX y. P y" ..
+qed
+
+text {*
+ Technically, \isakeyword{obtain} is similar to \isakeyword{fix} and
+ \isakeyword{assume} together with a soundness proof of the
+ elimination involved. Thus it behaves similar to any other forward
+ proof element. Also note that due to the nature of general
+ existence reasoning involved here, any result exported from the
+ context of an \isakeyword{obtain} statement may \emph{not} refer to
+ the parameters introduced there.
+*}
+
+
+
+subsubsection {* Deriving rules in Isabelle *}
+
+text {*
+ We derive the conjunction elimination rule from the corresponding
+ projections. The proof is quite straight-forward, since
+ Isabelle/Isar supports non-atomic goals and assumptions fully
+ transparently.
+*}
+
+theorem conjE: "A & B ==> (A ==> B ==> C) ==> C"
+proof -
+ assume "A & B"
+ assume r: "A ==> B ==> C"
+ show C
+ proof (rule r)
+ show A by (rule conjunct1) fact
+ show B by (rule conjunct2) fact
+ qed
+qed
+
+end
--- a/src/HOL/Isar_examples/Cantor.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Cantor.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,11 +1,12 @@
(* Title: HOL/Isar_examples/Cantor.thy
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
*)
header {* Cantor's Theorem *}
-theory Cantor imports Main begin
+theory Cantor
+imports Main
+begin
text_raw {*
\footnote{This is an Isar version of the final example of the
--- a/src/HOL/Isar_examples/Drinker.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Drinker.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,11 +1,12 @@
(* Title: HOL/Isar_examples/Drinker.thy
- ID: $Id$
Author: Makarius
*)
header {* The Drinker's Principle *}
-theory Drinker imports Main begin
+theory Drinker
+imports Main
+begin
text {*
Here is another example of classical reasoning: the Drinker's
--- a/src/HOL/Isar_examples/ExprCompiler.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,230 +0,0 @@
-(* Title: HOL/Isar_examples/ExprCompiler.thy
- ID: $Id$
- Author: Markus Wenzel, TU Muenchen
-
-Correctness of a simple expression/stack-machine compiler.
-*)
-
-header {* Correctness of a simple expression compiler *}
-
-theory ExprCompiler imports Main begin
-
-text {*
- This is a (rather trivial) example of program verification. We model
- a compiler for translating expressions to stack machine instructions,
- and prove its correctness wrt.\ some evaluation semantics.
-*}
-
-
-subsection {* Binary operations *}
-
-text {*
- Binary operations are just functions over some type of values. This
- is both for abstract syntax and semantics, i.e.\ we use a ``shallow
- embedding'' here.
-*}
-
-types
- 'val binop = "'val => 'val => 'val"
-
-
-subsection {* Expressions *}
-
-text {*
- The language of expressions is defined as an inductive type,
- consisting of variables, constants, and binary operations on
- expressions.
-*}
-
-datatype ('adr, 'val) expr =
- Variable 'adr |
- Constant 'val |
- Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr"
-
-text {*
- Evaluation (wrt.\ some environment of variable assignments) is
- defined by primitive recursion over the structure of expressions.
-*}
-
-consts
- eval :: "('adr, 'val) expr => ('adr => 'val) => 'val"
-
-primrec
- "eval (Variable x) env = env x"
- "eval (Constant c) env = c"
- "eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)"
-
-
-subsection {* Machine *}
-
-text {*
- Next we model a simple stack machine, with three instructions.
-*}
-
-datatype ('adr, 'val) instr =
- Const 'val |
- Load 'adr |
- Apply "'val binop"
-
-text {*
- Execution of a list of stack machine instructions is easily defined
- as follows.
-*}
-
-consts
- exec :: "(('adr, 'val) instr) list
- => 'val list => ('adr => 'val) => 'val list"
-
-primrec
- "exec [] stack env = stack"
- "exec (instr # instrs) stack env =
- (case instr of
- Const c => exec instrs (c # stack) env
- | Load x => exec instrs (env x # stack) env
- | Apply f => exec instrs (f (hd stack) (hd (tl stack))
- # (tl (tl stack))) env)"
-
-constdefs
- execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val"
- "execute instrs env == hd (exec instrs [] env)"
-
-
-subsection {* Compiler *}
-
-text {*
- We are ready to define the compilation function of expressions to
- lists of stack machine instructions.
-*}
-
-consts
- compile :: "('adr, 'val) expr => (('adr, 'val) instr) list"
-
-primrec
- "compile (Variable x) = [Load x]"
- "compile (Constant c) = [Const c]"
- "compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]"
-
-
-text {*
- The main result of this development is the correctness theorem for
- $\idt{compile}$. We first establish a lemma about $\idt{exec}$ and
- list append.
-*}
-
-lemma exec_append:
- "exec (xs @ ys) stack env =
- exec ys (exec xs stack env) env"
-proof (induct xs arbitrary: stack)
- case Nil
- show ?case by simp
-next
- case (Cons x xs)
- show ?case
- proof (induct x)
- case Const
- from Cons show ?case by simp
- next
- case Load
- from Cons show ?case by simp
- next
- case Apply
- from Cons show ?case by simp
- qed
-qed
-
-theorem correctness: "execute (compile e) env = eval e env"
-proof -
- have "\<And>stack. exec (compile e) stack env = eval e env # stack"
- proof (induct e)
- case Variable show ?case by simp
- next
- case Constant show ?case by simp
- next
- case Binop then show ?case by (simp add: exec_append)
- qed
- then show ?thesis by (simp add: execute_def)
-qed
-
-
-text {*
- \bigskip In the proofs above, the \name{simp} method does quite a lot
- of work behind the scenes (mostly ``functional program execution'').
- Subsequently, the same reasoning is elaborated in detail --- at most
- one recursive function definition is used at a time. Thus we get a
- better idea of what is actually going on.
-*}
-
-lemma exec_append':
- "exec (xs @ ys) stack env = exec ys (exec xs stack env) env"
-proof (induct xs arbitrary: stack)
- case (Nil s)
- have "exec ([] @ ys) s env = exec ys s env" by simp
- also have "... = exec ys (exec [] s env) env" by simp
- finally show ?case .
-next
- case (Cons x xs s)
- show ?case
- proof (induct x)
- case (Const val)
- have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env"
- by simp
- also have "... = exec (xs @ ys) (val # s) env" by simp
- also from Cons have "... = exec ys (exec xs (val # s) env) env" .
- also have "... = exec ys (exec (Const val # xs) s env) env" by simp
- finally show ?case .
- next
- case (Load adr)
- from Cons show ?case by simp -- {* same as above *}
- next
- case (Apply fn)
- have "exec ((Apply fn # xs) @ ys) s env =
- exec (Apply fn # xs @ ys) s env" by simp
- also have "... =
- exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp
- also from Cons have "... =
- exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" .
- also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp
- finally show ?case .
- qed
-qed
-
-theorem correctness': "execute (compile e) env = eval e env"
-proof -
- have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack"
- proof (induct e)
- case (Variable adr s)
- have "exec (compile (Variable adr)) s env = exec [Load adr] s env"
- by simp
- also have "... = env adr # s" by simp
- also have "env adr = eval (Variable adr) env" by simp
- finally show ?case .
- next
- case (Constant val s)
- show ?case by simp -- {* same as above *}
- next
- case (Binop fn e1 e2 s)
- have "exec (compile (Binop fn e1 e2)) s env =
- exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp
- also have "... = exec [Apply fn]
- (exec (compile e1) (exec (compile e2) s env) env) env"
- by (simp only: exec_append)
- also have "exec (compile e2) s env = eval e2 env # s" by fact
- also have "exec (compile e1) ... env = eval e1 env # ..." by fact
- also have "exec [Apply fn] ... env =
- fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp
- also have "... = fn (eval e1 env) (eval e2 env) # s" by simp
- also have "fn (eval e1 env) (eval e2 env) =
- eval (Binop fn e1 e2) env"
- by simp
- finally show ?case .
- qed
-
- have "execute (compile e) env = hd (exec (compile e) [] env)"
- by (simp add: execute_def)
- also from exec_compile
- have "exec (compile e) [] env = [eval e env]" .
- also have "hd ... = eval e env" by simp
- finally show ?thesis .
-qed
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Expr_Compiler.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,231 @@
+(* Title: HOL/Isar_examples/Expr_Compiler.thy
+ Author: Markus Wenzel, TU Muenchen
+
+Correctness of a simple expression/stack-machine compiler.
+*)
+
+header {* Correctness of a simple expression compiler *}
+
+theory Expr_Compiler
+imports Main
+begin
+
+text {*
+ This is a (rather trivial) example of program verification. We model
+ a compiler for translating expressions to stack machine instructions,
+ and prove its correctness wrt.\ some evaluation semantics.
+*}
+
+
+subsection {* Binary operations *}
+
+text {*
+ Binary operations are just functions over some type of values. This
+ is both for abstract syntax and semantics, i.e.\ we use a ``shallow
+ embedding'' here.
+*}
+
+types
+ 'val binop = "'val => 'val => 'val"
+
+
+subsection {* Expressions *}
+
+text {*
+ The language of expressions is defined as an inductive type,
+ consisting of variables, constants, and binary operations on
+ expressions.
+*}
+
+datatype ('adr, 'val) expr =
+ Variable 'adr |
+ Constant 'val |
+ Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr"
+
+text {*
+ Evaluation (wrt.\ some environment of variable assignments) is
+ defined by primitive recursion over the structure of expressions.
+*}
+
+consts
+ eval :: "('adr, 'val) expr => ('adr => 'val) => 'val"
+
+primrec
+ "eval (Variable x) env = env x"
+ "eval (Constant c) env = c"
+ "eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)"
+
+
+subsection {* Machine *}
+
+text {*
+ Next we model a simple stack machine, with three instructions.
+*}
+
+datatype ('adr, 'val) instr =
+ Const 'val |
+ Load 'adr |
+ Apply "'val binop"
+
+text {*
+ Execution of a list of stack machine instructions is easily defined
+ as follows.
+*}
+
+consts
+ exec :: "(('adr, 'val) instr) list
+ => 'val list => ('adr => 'val) => 'val list"
+
+primrec
+ "exec [] stack env = stack"
+ "exec (instr # instrs) stack env =
+ (case instr of
+ Const c => exec instrs (c # stack) env
+ | Load x => exec instrs (env x # stack) env
+ | Apply f => exec instrs (f (hd stack) (hd (tl stack))
+ # (tl (tl stack))) env)"
+
+constdefs
+ execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val"
+ "execute instrs env == hd (exec instrs [] env)"
+
+
+subsection {* Compiler *}
+
+text {*
+ We are ready to define the compilation function of expressions to
+ lists of stack machine instructions.
+*}
+
+consts
+ compile :: "('adr, 'val) expr => (('adr, 'val) instr) list"
+
+primrec
+ "compile (Variable x) = [Load x]"
+ "compile (Constant c) = [Const c]"
+ "compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]"
+
+
+text {*
+ The main result of this development is the correctness theorem for
+ $\idt{compile}$. We first establish a lemma about $\idt{exec}$ and
+ list append.
+*}
+
+lemma exec_append:
+ "exec (xs @ ys) stack env =
+ exec ys (exec xs stack env) env"
+proof (induct xs arbitrary: stack)
+ case Nil
+ show ?case by simp
+next
+ case (Cons x xs)
+ show ?case
+ proof (induct x)
+ case Const
+ from Cons show ?case by simp
+ next
+ case Load
+ from Cons show ?case by simp
+ next
+ case Apply
+ from Cons show ?case by simp
+ qed
+qed
+
+theorem correctness: "execute (compile e) env = eval e env"
+proof -
+ have "\<And>stack. exec (compile e) stack env = eval e env # stack"
+ proof (induct e)
+ case Variable show ?case by simp
+ next
+ case Constant show ?case by simp
+ next
+ case Binop then show ?case by (simp add: exec_append)
+ qed
+ then show ?thesis by (simp add: execute_def)
+qed
+
+
+text {*
+ \bigskip In the proofs above, the \name{simp} method does quite a lot
+ of work behind the scenes (mostly ``functional program execution'').
+ Subsequently, the same reasoning is elaborated in detail --- at most
+ one recursive function definition is used at a time. Thus we get a
+ better idea of what is actually going on.
+*}
+
+lemma exec_append':
+ "exec (xs @ ys) stack env = exec ys (exec xs stack env) env"
+proof (induct xs arbitrary: stack)
+ case (Nil s)
+ have "exec ([] @ ys) s env = exec ys s env" by simp
+ also have "... = exec ys (exec [] s env) env" by simp
+ finally show ?case .
+next
+ case (Cons x xs s)
+ show ?case
+ proof (induct x)
+ case (Const val)
+ have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env"
+ by simp
+ also have "... = exec (xs @ ys) (val # s) env" by simp
+ also from Cons have "... = exec ys (exec xs (val # s) env) env" .
+ also have "... = exec ys (exec (Const val # xs) s env) env" by simp
+ finally show ?case .
+ next
+ case (Load adr)
+ from Cons show ?case by simp -- {* same as above *}
+ next
+ case (Apply fn)
+ have "exec ((Apply fn # xs) @ ys) s env =
+ exec (Apply fn # xs @ ys) s env" by simp
+ also have "... =
+ exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp
+ also from Cons have "... =
+ exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" .
+ also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp
+ finally show ?case .
+ qed
+qed
+
+theorem correctness': "execute (compile e) env = eval e env"
+proof -
+ have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack"
+ proof (induct e)
+ case (Variable adr s)
+ have "exec (compile (Variable adr)) s env = exec [Load adr] s env"
+ by simp
+ also have "... = env adr # s" by simp
+ also have "env adr = eval (Variable adr) env" by simp
+ finally show ?case .
+ next
+ case (Constant val s)
+ show ?case by simp -- {* same as above *}
+ next
+ case (Binop fn e1 e2 s)
+ have "exec (compile (Binop fn e1 e2)) s env =
+ exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp
+ also have "... = exec [Apply fn]
+ (exec (compile e1) (exec (compile e2) s env) env) env"
+ by (simp only: exec_append)
+ also have "exec (compile e2) s env = eval e2 env # s" by fact
+ also have "exec (compile e1) ... env = eval e1 env # ..." by fact
+ also have "exec [Apply fn] ... env =
+ fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp
+ also have "... = fn (eval e1 env) (eval e2 env) # s" by simp
+ also have "fn (eval e1 env) (eval e2 env) =
+ eval (Binop fn e1 e2) env"
+ by simp
+ finally show ?case .
+ qed
+
+ have "execute (compile e) env = hd (exec (compile e) [] env)"
+ by (simp add: execute_def)
+ also from exec_compile
+ have "exec (compile e) [] env = [eval e env]" .
+ also have "hd ... = eval e env" by simp
+ finally show ?thesis .
+qed
+
+end
--- a/src/HOL/Isar_examples/Fibonacci.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Fibonacci.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Isar_examples/Fibonacci.thy
- ID: $Id$
Author: Gertrud Bauer
Copyright 1999 Technische Universitaet Muenchen
--- a/src/HOL/Isar_examples/Group.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Group.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,11 +1,12 @@
(* Title: HOL/Isar_examples/Group.thy
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
*)
header {* Basic group theory *}
-theory Group imports Main begin
+theory Group
+imports Main
+begin
subsection {* Groups and calculational reasoning *}
--- a/src/HOL/Isar_examples/Hoare.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Hoare.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Isar_examples/Hoare.thy
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
A formulation of Hoare logic suitable for Isar.
@@ -7,8 +6,10 @@
header {* Hoare Logic *}
-theory Hoare imports Main
-uses ("~~/src/HOL/Hoare/hoare_tac.ML") begin
+theory Hoare
+imports Main
+uses ("~~/src/HOL/Hoare/hoare_tac.ML")
+begin
subsection {* Abstract syntax and semantics *}
@@ -260,7 +261,7 @@
| bexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/Isar_examples/HoareEx.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,328 +0,0 @@
-
-header {* Using Hoare Logic *}
-
-theory HoareEx imports Hoare begin
-
-subsection {* State spaces *}
-
-text {*
- First of all we provide a store of program variables that
- occur in any of the programs considered later. Slightly unexpected
- things may happen when attempting to work with undeclared variables.
-*}
-
-record vars =
- I :: nat
- M :: nat
- N :: nat
- S :: nat
-
-text {*
- While all of our variables happen to have the same type, nothing
- would prevent us from working with many-sorted programs as well, or
- even polymorphic ones. Also note that Isabelle/HOL's extensible
- record types even provides simple means to extend the state space
- later.
-*}
-
-
-subsection {* Basic examples *}
-
-text {*
- We look at few trivialities involving assignment and sequential
- composition, in order to get an idea of how to work with our
- formulation of Hoare Logic.
-*}
-
-text {*
- Using the basic \name{assign} rule directly is a bit cumbersome.
-*}
-
-lemma
- "|- .{\<acute>(N_update (\<lambda>_. (2 * \<acute>N))) : .{\<acute>N = 10}.}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
- by (rule assign)
-
-text {*
- Certainly we want the state modification already done, e.g.\ by
- simplification. The \name{hoare} method performs the basic state
- update for us; we may apply the Simplifier afterwards to achieve
- ``obvious'' consequences as well.
-*}
-
-lemma "|- .{True}. \<acute>N := 10 .{\<acute>N = 10}."
- by hoare
-
-lemma "|- .{2 * \<acute>N = 10}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
- by hoare
-
-lemma "|- .{\<acute>N = 5}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
- by hoare simp
-
-lemma "|- .{\<acute>N + 1 = a + 1}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
- by hoare
-
-lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
- by hoare simp
-
-lemma "|- .{a = a & b = b}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
- by hoare
-
-lemma "|- .{True}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
- by hoare simp
-
-lemma
-"|- .{\<acute>M = a & \<acute>N = b}.
- \<acute>I := \<acute>M; \<acute>M := \<acute>N; \<acute>N := \<acute>I
- .{\<acute>M = b & \<acute>N = a}."
- by hoare simp
-
-text {*
- It is important to note that statements like the following one can
- only be proven for each individual program variable. Due to the
- extra-logical nature of record fields, we cannot formulate a theorem
- relating record selectors and updates schematically.
-*}
-
-lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N .{\<acute>N = a}."
- by hoare
-
-lemma "|- .{\<acute>x = a}. \<acute>x := \<acute>x .{\<acute>x = a}."
- oops
-
-lemma
- "Valid {s. x s = a} (Basic (\<lambda>s. x_update (x s) s)) {s. x s = n}"
- -- {* same statement without concrete syntax *}
- oops
-
-
-text {*
- In the following assignments we make use of the consequence rule in
- order to achieve the intended precondition. Certainly, the
- \name{hoare} method is able to handle this case, too.
-*}
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-proof -
- have ".{\<acute>M = \<acute>N}. <= .{\<acute>M + 1 ~= \<acute>N}."
- by auto
- also have "|- ... \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
- by hoare
- finally show ?thesis .
-qed
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-proof -
- have "!!m n::nat. m = n --> m + 1 ~= n"
- -- {* inclusion of assertions expressed in ``pure'' logic, *}
- -- {* without mentioning the state space *}
- by simp
- also have "|- .{\<acute>M + 1 ~= \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
- by hoare
- finally show ?thesis .
-qed
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
- by hoare simp
-
-
-subsection {* Multiplication by addition *}
-
-text {*
- We now do some basic examples of actual \texttt{WHILE} programs.
- This one is a loop for calculating the product of two natural
- numbers, by iterated addition. We first give detailed structured
- proof based on single-step Hoare rules.
-*}
-
-lemma
- "|- .{\<acute>M = 0 & \<acute>S = 0}.
- WHILE \<acute>M ~= a
- DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
- .{\<acute>S = a * b}."
-proof -
- let "|- _ ?while _" = ?thesis
- let ".{\<acute>?inv}." = ".{\<acute>S = \<acute>M * b}."
-
- have ".{\<acute>M = 0 & \<acute>S = 0}. <= .{\<acute>?inv}." by auto
- also have "|- ... ?while .{\<acute>?inv & ~ (\<acute>M ~= a)}."
- proof
- let ?c = "\<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1"
- have ".{\<acute>?inv & \<acute>M ~= a}. <= .{\<acute>S + b = (\<acute>M + 1) * b}."
- by auto
- also have "|- ... ?c .{\<acute>?inv}." by hoare
- finally show "|- .{\<acute>?inv & \<acute>M ~= a}. ?c .{\<acute>?inv}." .
- qed
- also have "... <= .{\<acute>S = a * b}." by auto
- finally show ?thesis .
-qed
-
-text {*
- The subsequent version of the proof applies the \name{hoare} method
- to reduce the Hoare statement to a purely logical problem that can be
- solved fully automatically. Note that we have to specify the
- \texttt{WHILE} loop invariant in the original statement.
-*}
-
-lemma
- "|- .{\<acute>M = 0 & \<acute>S = 0}.
- WHILE \<acute>M ~= a
- INV .{\<acute>S = \<acute>M * b}.
- DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
- .{\<acute>S = a * b}."
- by hoare auto
-
-
-subsection {* Summing natural numbers *}
-
-text {*
- We verify an imperative program to sum natural numbers up to a given
- limit. First some functional definition for proper specification of
- the problem.
-*}
-
-text {*
- The following proof is quite explicit in the individual steps taken,
- with the \name{hoare} method only applied locally to take care of
- assignment and sequential composition. Note that we express
- intermediate proof obligation in pure logic, without referring to the
- state space.
-*}
-
-declare atLeast0LessThan[symmetric,simp]
-
-theorem
- "|- .{True}.
- \<acute>S := 0; \<acute>I := 1;
- WHILE \<acute>I ~= n
- DO
- \<acute>S := \<acute>S + \<acute>I;
- \<acute>I := \<acute>I + 1
- OD
- .{\<acute>S = (SUM j<n. j)}."
- (is "|- _ (_; ?while) _")
-proof -
- let ?sum = "\<lambda>k::nat. SUM j<k. j"
- let ?inv = "\<lambda>s i::nat. s = ?sum i"
-
- have "|- .{True}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
- proof -
- have "True --> 0 = ?sum 1"
- by simp
- also have "|- .{...}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
- by hoare
- finally show ?thesis .
- qed
- also have "|- ... ?while .{?inv \<acute>S \<acute>I & ~ \<acute>I ~= n}."
- proof
- let ?body = "\<acute>S := \<acute>S + \<acute>I; \<acute>I := \<acute>I + 1"
- have "!!s i. ?inv s i & i ~= n --> ?inv (s + i) (i + 1)"
- by simp
- also have "|- .{\<acute>S + \<acute>I = ?sum (\<acute>I + 1)}. ?body .{?inv \<acute>S \<acute>I}."
- by hoare
- finally show "|- .{?inv \<acute>S \<acute>I & \<acute>I ~= n}. ?body .{?inv \<acute>S \<acute>I}." .
- qed
- also have "!!s i. s = ?sum i & ~ i ~= n --> s = ?sum n"
- by simp
- finally show ?thesis .
-qed
-
-text {*
- The next version uses the \name{hoare} method, while still explaining
- the resulting proof obligations in an abstract, structured manner.
-*}
-
-theorem
- "|- .{True}.
- \<acute>S := 0; \<acute>I := 1;
- WHILE \<acute>I ~= n
- INV .{\<acute>S = (SUM j<\<acute>I. j)}.
- DO
- \<acute>S := \<acute>S + \<acute>I;
- \<acute>I := \<acute>I + 1
- OD
- .{\<acute>S = (SUM j<n. j)}."
-proof -
- let ?sum = "\<lambda>k::nat. SUM j<k. j"
- let ?inv = "\<lambda>s i::nat. s = ?sum i"
-
- show ?thesis
- proof hoare
- show "?inv 0 1" by simp
- next
- fix s i assume "?inv s i & i ~= n"
- thus "?inv (s + i) (i + 1)" by simp
- next
- fix s i assume "?inv s i & ~ i ~= n"
- thus "s = ?sum n" by simp
- qed
-qed
-
-text {*
- Certainly, this proof may be done fully automatic as well, provided
- that the invariant is given beforehand.
-*}
-
-theorem
- "|- .{True}.
- \<acute>S := 0; \<acute>I := 1;
- WHILE \<acute>I ~= n
- INV .{\<acute>S = (SUM j<\<acute>I. j)}.
- DO
- \<acute>S := \<acute>S + \<acute>I;
- \<acute>I := \<acute>I + 1
- OD
- .{\<acute>S = (SUM j<n. j)}."
- by hoare auto
-
-
-subsection{* Time *}
-
-text{*
- A simple embedding of time in Hoare logic: function @{text timeit}
- inserts an extra variable to keep track of the elapsed time.
-*}
-
-record tstate = time :: nat
-
-types 'a time = "\<lparr>time :: nat, \<dots> :: 'a\<rparr>"
-
-consts timeit :: "'a time com \<Rightarrow> 'a time com"
-primrec
- "timeit (Basic f) = (Basic f; Basic(\<lambda>s. s\<lparr>time := Suc (time s)\<rparr>))"
- "timeit (c1; c2) = (timeit c1; timeit c2)"
- "timeit (Cond b c1 c2) = Cond b (timeit c1) (timeit c2)"
- "timeit (While b iv c) = While b iv (timeit c)"
-
-record tvars = tstate +
- I :: nat
- J :: nat
-
-lemma lem: "(0::nat) < n \<Longrightarrow> n + n \<le> Suc (n * n)"
- by (induct n) simp_all
-
-lemma "|- .{i = \<acute>I & \<acute>time = 0}.
- timeit(
- WHILE \<acute>I \<noteq> 0
- INV .{2*\<acute>time + \<acute>I*\<acute>I + 5*\<acute>I = i*i + 5*i}.
- DO
- \<acute>J := \<acute>I;
- WHILE \<acute>J \<noteq> 0
- INV .{0 < \<acute>I & 2*\<acute>time + \<acute>I*\<acute>I + 3*\<acute>I + 2*\<acute>J - 2 = i*i + 5*i}.
- DO \<acute>J := \<acute>J - 1 OD;
- \<acute>I := \<acute>I - 1
- OD
- ) .{2*\<acute>time = i*i + 5*i}."
- apply simp
- apply hoare
- apply simp
- apply clarsimp
- apply clarsimp
- apply arith
- prefer 2
- apply clarsimp
- apply (clarsimp simp: nat_distrib)
- apply (frule lem)
- apply arith
- done
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Hoare_Ex.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,329 @@
+header {* Using Hoare Logic *}
+
+theory Hoare_Ex
+imports Hoare
+begin
+
+subsection {* State spaces *}
+
+text {*
+ First of all we provide a store of program variables that
+ occur in any of the programs considered later. Slightly unexpected
+ things may happen when attempting to work with undeclared variables.
+*}
+
+record vars =
+ I :: nat
+ M :: nat
+ N :: nat
+ S :: nat
+
+text {*
+ While all of our variables happen to have the same type, nothing
+ would prevent us from working with many-sorted programs as well, or
+ even polymorphic ones. Also note that Isabelle/HOL's extensible
+ record types even provides simple means to extend the state space
+ later.
+*}
+
+
+subsection {* Basic examples *}
+
+text {*
+ We look at few trivialities involving assignment and sequential
+ composition, in order to get an idea of how to work with our
+ formulation of Hoare Logic.
+*}
+
+text {*
+ Using the basic \name{assign} rule directly is a bit cumbersome.
+*}
+
+lemma
+ "|- .{\<acute>(N_update (\<lambda>_. (2 * \<acute>N))) : .{\<acute>N = 10}.}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+ by (rule assign)
+
+text {*
+ Certainly we want the state modification already done, e.g.\ by
+ simplification. The \name{hoare} method performs the basic state
+ update for us; we may apply the Simplifier afterwards to achieve
+ ``obvious'' consequences as well.
+*}
+
+lemma "|- .{True}. \<acute>N := 10 .{\<acute>N = 10}."
+ by hoare
+
+lemma "|- .{2 * \<acute>N = 10}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+ by hoare
+
+lemma "|- .{\<acute>N = 5}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+ by hoare simp
+
+lemma "|- .{\<acute>N + 1 = a + 1}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
+ by hoare
+
+lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
+ by hoare simp
+
+lemma "|- .{a = a & b = b}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
+ by hoare
+
+lemma "|- .{True}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
+ by hoare simp
+
+lemma
+"|- .{\<acute>M = a & \<acute>N = b}.
+ \<acute>I := \<acute>M; \<acute>M := \<acute>N; \<acute>N := \<acute>I
+ .{\<acute>M = b & \<acute>N = a}."
+ by hoare simp
+
+text {*
+ It is important to note that statements like the following one can
+ only be proven for each individual program variable. Due to the
+ extra-logical nature of record fields, we cannot formulate a theorem
+ relating record selectors and updates schematically.
+*}
+
+lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N .{\<acute>N = a}."
+ by hoare
+
+lemma "|- .{\<acute>x = a}. \<acute>x := \<acute>x .{\<acute>x = a}."
+ oops
+
+lemma
+ "Valid {s. x s = a} (Basic (\<lambda>s. x_update (x s) s)) {s. x s = n}"
+ -- {* same statement without concrete syntax *}
+ oops
+
+
+text {*
+ In the following assignments we make use of the consequence rule in
+ order to achieve the intended precondition. Certainly, the
+ \name{hoare} method is able to handle this case, too.
+*}
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+proof -
+ have ".{\<acute>M = \<acute>N}. <= .{\<acute>M + 1 ~= \<acute>N}."
+ by auto
+ also have "|- ... \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+ by hoare
+ finally show ?thesis .
+qed
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+proof -
+ have "!!m n::nat. m = n --> m + 1 ~= n"
+ -- {* inclusion of assertions expressed in ``pure'' logic, *}
+ -- {* without mentioning the state space *}
+ by simp
+ also have "|- .{\<acute>M + 1 ~= \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+ by hoare
+ finally show ?thesis .
+qed
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+ by hoare simp
+
+
+subsection {* Multiplication by addition *}
+
+text {*
+ We now do some basic examples of actual \texttt{WHILE} programs.
+ This one is a loop for calculating the product of two natural
+ numbers, by iterated addition. We first give detailed structured
+ proof based on single-step Hoare rules.
+*}
+
+lemma
+ "|- .{\<acute>M = 0 & \<acute>S = 0}.
+ WHILE \<acute>M ~= a
+ DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
+ .{\<acute>S = a * b}."
+proof -
+ let "|- _ ?while _" = ?thesis
+ let ".{\<acute>?inv}." = ".{\<acute>S = \<acute>M * b}."
+
+ have ".{\<acute>M = 0 & \<acute>S = 0}. <= .{\<acute>?inv}." by auto
+ also have "|- ... ?while .{\<acute>?inv & ~ (\<acute>M ~= a)}."
+ proof
+ let ?c = "\<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1"
+ have ".{\<acute>?inv & \<acute>M ~= a}. <= .{\<acute>S + b = (\<acute>M + 1) * b}."
+ by auto
+ also have "|- ... ?c .{\<acute>?inv}." by hoare
+ finally show "|- .{\<acute>?inv & \<acute>M ~= a}. ?c .{\<acute>?inv}." .
+ qed
+ also have "... <= .{\<acute>S = a * b}." by auto
+ finally show ?thesis .
+qed
+
+text {*
+ The subsequent version of the proof applies the \name{hoare} method
+ to reduce the Hoare statement to a purely logical problem that can be
+ solved fully automatically. Note that we have to specify the
+ \texttt{WHILE} loop invariant in the original statement.
+*}
+
+lemma
+ "|- .{\<acute>M = 0 & \<acute>S = 0}.
+ WHILE \<acute>M ~= a
+ INV .{\<acute>S = \<acute>M * b}.
+ DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
+ .{\<acute>S = a * b}."
+ by hoare auto
+
+
+subsection {* Summing natural numbers *}
+
+text {*
+ We verify an imperative program to sum natural numbers up to a given
+ limit. First some functional definition for proper specification of
+ the problem.
+*}
+
+text {*
+ The following proof is quite explicit in the individual steps taken,
+ with the \name{hoare} method only applied locally to take care of
+ assignment and sequential composition. Note that we express
+ intermediate proof obligation in pure logic, without referring to the
+ state space.
+*}
+
+declare atLeast0LessThan[symmetric,simp]
+
+theorem
+ "|- .{True}.
+ \<acute>S := 0; \<acute>I := 1;
+ WHILE \<acute>I ~= n
+ DO
+ \<acute>S := \<acute>S + \<acute>I;
+ \<acute>I := \<acute>I + 1
+ OD
+ .{\<acute>S = (SUM j<n. j)}."
+ (is "|- _ (_; ?while) _")
+proof -
+ let ?sum = "\<lambda>k::nat. SUM j<k. j"
+ let ?inv = "\<lambda>s i::nat. s = ?sum i"
+
+ have "|- .{True}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
+ proof -
+ have "True --> 0 = ?sum 1"
+ by simp
+ also have "|- .{...}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
+ by hoare
+ finally show ?thesis .
+ qed
+ also have "|- ... ?while .{?inv \<acute>S \<acute>I & ~ \<acute>I ~= n}."
+ proof
+ let ?body = "\<acute>S := \<acute>S + \<acute>I; \<acute>I := \<acute>I + 1"
+ have "!!s i. ?inv s i & i ~= n --> ?inv (s + i) (i + 1)"
+ by simp
+ also have "|- .{\<acute>S + \<acute>I = ?sum (\<acute>I + 1)}. ?body .{?inv \<acute>S \<acute>I}."
+ by hoare
+ finally show "|- .{?inv \<acute>S \<acute>I & \<acute>I ~= n}. ?body .{?inv \<acute>S \<acute>I}." .
+ qed
+ also have "!!s i. s = ?sum i & ~ i ~= n --> s = ?sum n"
+ by simp
+ finally show ?thesis .
+qed
+
+text {*
+ The next version uses the \name{hoare} method, while still explaining
+ the resulting proof obligations in an abstract, structured manner.
+*}
+
+theorem
+ "|- .{True}.
+ \<acute>S := 0; \<acute>I := 1;
+ WHILE \<acute>I ~= n
+ INV .{\<acute>S = (SUM j<\<acute>I. j)}.
+ DO
+ \<acute>S := \<acute>S + \<acute>I;
+ \<acute>I := \<acute>I + 1
+ OD
+ .{\<acute>S = (SUM j<n. j)}."
+proof -
+ let ?sum = "\<lambda>k::nat. SUM j<k. j"
+ let ?inv = "\<lambda>s i::nat. s = ?sum i"
+
+ show ?thesis
+ proof hoare
+ show "?inv 0 1" by simp
+ next
+ fix s i assume "?inv s i & i ~= n"
+ thus "?inv (s + i) (i + 1)" by simp
+ next
+ fix s i assume "?inv s i & ~ i ~= n"
+ thus "s = ?sum n" by simp
+ qed
+qed
+
+text {*
+ Certainly, this proof may be done fully automatic as well, provided
+ that the invariant is given beforehand.
+*}
+
+theorem
+ "|- .{True}.
+ \<acute>S := 0; \<acute>I := 1;
+ WHILE \<acute>I ~= n
+ INV .{\<acute>S = (SUM j<\<acute>I. j)}.
+ DO
+ \<acute>S := \<acute>S + \<acute>I;
+ \<acute>I := \<acute>I + 1
+ OD
+ .{\<acute>S = (SUM j<n. j)}."
+ by hoare auto
+
+
+subsection{* Time *}
+
+text{*
+ A simple embedding of time in Hoare logic: function @{text timeit}
+ inserts an extra variable to keep track of the elapsed time.
+*}
+
+record tstate = time :: nat
+
+types 'a time = "\<lparr>time :: nat, \<dots> :: 'a\<rparr>"
+
+consts timeit :: "'a time com \<Rightarrow> 'a time com"
+primrec
+ "timeit (Basic f) = (Basic f; Basic(\<lambda>s. s\<lparr>time := Suc (time s)\<rparr>))"
+ "timeit (c1; c2) = (timeit c1; timeit c2)"
+ "timeit (Cond b c1 c2) = Cond b (timeit c1) (timeit c2)"
+ "timeit (While b iv c) = While b iv (timeit c)"
+
+record tvars = tstate +
+ I :: nat
+ J :: nat
+
+lemma lem: "(0::nat) < n \<Longrightarrow> n + n \<le> Suc (n * n)"
+ by (induct n) simp_all
+
+lemma "|- .{i = \<acute>I & \<acute>time = 0}.
+ timeit(
+ WHILE \<acute>I \<noteq> 0
+ INV .{2*\<acute>time + \<acute>I*\<acute>I + 5*\<acute>I = i*i + 5*i}.
+ DO
+ \<acute>J := \<acute>I;
+ WHILE \<acute>J \<noteq> 0
+ INV .{0 < \<acute>I & 2*\<acute>time + \<acute>I*\<acute>I + 3*\<acute>I + 2*\<acute>J - 2 = i*i + 5*i}.
+ DO \<acute>J := \<acute>J - 1 OD;
+ \<acute>I := \<acute>I - 1
+ OD
+ ) .{2*\<acute>time = i*i + 5*i}."
+ apply simp
+ apply hoare
+ apply simp
+ apply clarsimp
+ apply clarsimp
+ apply arith
+ prefer 2
+ apply clarsimp
+ apply (clarsimp simp: nat_distrib)
+ apply (frule lem)
+ apply arith
+ done
+
+end
--- a/src/HOL/Isar_examples/KnasterTarski.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-(* Title: HOL/Isar_examples/KnasterTarski.thy
- Author: Markus Wenzel, TU Muenchen
-
-Typical textbook proof example.
-*)
-
-header {* Textbook-style reasoning: the Knaster-Tarski Theorem *}
-
-theory KnasterTarski
-imports Main Lattice_Syntax
-begin
-
-
-subsection {* Prose version *}
-
-text {*
- According to the textbook \cite[pages 93--94]{davey-priestley}, the
- Knaster-Tarski fixpoint theorem is as follows.\footnote{We have
- dualized the argument, and tuned the notation a little bit.}
-
- \textbf{The Knaster-Tarski Fixpoint Theorem.} Let @{text L} be a
- complete lattice and @{text "f: L \<rightarrow> L"} an order-preserving map.
- Then @{text "\<Sqinter>{x \<in> L | f(x) \<le> x}"} is a fixpoint of @{text f}.
-
- \textbf{Proof.} Let @{text "H = {x \<in> L | f(x) \<le> x}"} and @{text "a =
- \<Sqinter>H"}. For all @{text "x \<in> H"} we have @{text "a \<le> x"}, so @{text
- "f(a) \<le> f(x) \<le> x"}. Thus @{text "f(a)"} is a lower bound of @{text
- H}, whence @{text "f(a) \<le> a"}. We now use this inequality to prove
- the reverse one (!) and thereby complete the proof that @{text a} is
- a fixpoint. Since @{text f} is order-preserving, @{text "f(f(a)) \<le>
- f(a)"}. This says @{text "f(a) \<in> H"}, so @{text "a \<le> f(a)"}.
-*}
-
-
-subsection {* Formal versions *}
-
-text {*
- The Isar proof below closely follows the original presentation.
- Virtually all of the prose narration has been rephrased in terms of
- formal Isar language elements. Just as many textbook-style proofs,
- there is a strong bias towards forward proof, and several bends in
- the course of reasoning.
-*}
-
-theorem Knaster_Tarski:
- fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
- assumes "mono f"
- shows "\<exists>a. f a = a"
-proof
- let ?H = "{u. f u \<le> u}"
- let ?a = "\<Sqinter>?H"
- show "f ?a = ?a"
- proof -
- {
- fix x
- assume "x \<in> ?H"
- then have "?a \<le> x" by (rule Inf_lower)
- with `mono f` have "f ?a \<le> f x" ..
- also from `x \<in> ?H` have "\<dots> \<le> x" ..
- finally have "f ?a \<le> x" .
- }
- then have "f ?a \<le> ?a" by (rule Inf_greatest)
- {
- also presume "\<dots> \<le> f ?a"
- finally (order_antisym) show ?thesis .
- }
- from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
- then have "f ?a \<in> ?H" ..
- then show "?a \<le> f ?a" by (rule Inf_lower)
- qed
-qed
-
-text {*
- Above we have used several advanced Isar language elements, such as
- explicit block structure and weak assumptions. Thus we have
- mimicked the particular way of reasoning of the original text.
-
- In the subsequent version the order of reasoning is changed to
- achieve structured top-down decomposition of the problem at the
- outer level, while only the inner steps of reasoning are done in a
- forward manner. We are certainly more at ease here, requiring only
- the most basic features of the Isar language.
-*}
-
-theorem Knaster_Tarski':
- fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
- assumes "mono f"
- shows "\<exists>a. f a = a"
-proof
- let ?H = "{u. f u \<le> u}"
- let ?a = "\<Sqinter>?H"
- show "f ?a = ?a"
- proof (rule order_antisym)
- show "f ?a \<le> ?a"
- proof (rule Inf_greatest)
- fix x
- assume "x \<in> ?H"
- then have "?a \<le> x" by (rule Inf_lower)
- with `mono f` have "f ?a \<le> f x" ..
- also from `x \<in> ?H` have "\<dots> \<le> x" ..
- finally show "f ?a \<le> x" .
- qed
- show "?a \<le> f ?a"
- proof (rule Inf_lower)
- from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
- then show "f ?a \<in> ?H" ..
- qed
- qed
-qed
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Knaster_Tarski.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,111 @@
+(* Title: HOL/Isar_examples/Knaster_Tarski.thy
+ Author: Markus Wenzel, TU Muenchen
+
+Typical textbook proof example.
+*)
+
+header {* Textbook-style reasoning: the Knaster-Tarski Theorem *}
+
+theory Knaster_Tarski
+imports Main Lattice_Syntax
+begin
+
+
+subsection {* Prose version *}
+
+text {*
+ According to the textbook \cite[pages 93--94]{davey-priestley}, the
+ Knaster-Tarski fixpoint theorem is as follows.\footnote{We have
+ dualized the argument, and tuned the notation a little bit.}
+
+ \textbf{The Knaster-Tarski Fixpoint Theorem.} Let @{text L} be a
+ complete lattice and @{text "f: L \<rightarrow> L"} an order-preserving map.
+ Then @{text "\<Sqinter>{x \<in> L | f(x) \<le> x}"} is a fixpoint of @{text f}.
+
+ \textbf{Proof.} Let @{text "H = {x \<in> L | f(x) \<le> x}"} and @{text "a =
+ \<Sqinter>H"}. For all @{text "x \<in> H"} we have @{text "a \<le> x"}, so @{text
+ "f(a) \<le> f(x) \<le> x"}. Thus @{text "f(a)"} is a lower bound of @{text
+ H}, whence @{text "f(a) \<le> a"}. We now use this inequality to prove
+ the reverse one (!) and thereby complete the proof that @{text a} is
+ a fixpoint. Since @{text f} is order-preserving, @{text "f(f(a)) \<le>
+ f(a)"}. This says @{text "f(a) \<in> H"}, so @{text "a \<le> f(a)"}.
+*}
+
+
+subsection {* Formal versions *}
+
+text {*
+ The Isar proof below closely follows the original presentation.
+ Virtually all of the prose narration has been rephrased in terms of
+ formal Isar language elements. Just as many textbook-style proofs,
+ there is a strong bias towards forward proof, and several bends in
+ the course of reasoning.
+*}
+
+theorem Knaster_Tarski:
+ fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
+ assumes "mono f"
+ shows "\<exists>a. f a = a"
+proof
+ let ?H = "{u. f u \<le> u}"
+ let ?a = "\<Sqinter>?H"
+ show "f ?a = ?a"
+ proof -
+ {
+ fix x
+ assume "x \<in> ?H"
+ then have "?a \<le> x" by (rule Inf_lower)
+ with `mono f` have "f ?a \<le> f x" ..
+ also from `x \<in> ?H` have "\<dots> \<le> x" ..
+ finally have "f ?a \<le> x" .
+ }
+ then have "f ?a \<le> ?a" by (rule Inf_greatest)
+ {
+ also presume "\<dots> \<le> f ?a"
+ finally (order_antisym) show ?thesis .
+ }
+ from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
+ then have "f ?a \<in> ?H" ..
+ then show "?a \<le> f ?a" by (rule Inf_lower)
+ qed
+qed
+
+text {*
+ Above we have used several advanced Isar language elements, such as
+ explicit block structure and weak assumptions. Thus we have
+ mimicked the particular way of reasoning of the original text.
+
+ In the subsequent version the order of reasoning is changed to
+ achieve structured top-down decomposition of the problem at the
+ outer level, while only the inner steps of reasoning are done in a
+ forward manner. We are certainly more at ease here, requiring only
+ the most basic features of the Isar language.
+*}
+
+theorem Knaster_Tarski':
+ fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
+ assumes "mono f"
+ shows "\<exists>a. f a = a"
+proof
+ let ?H = "{u. f u \<le> u}"
+ let ?a = "\<Sqinter>?H"
+ show "f ?a = ?a"
+ proof (rule order_antisym)
+ show "f ?a \<le> ?a"
+ proof (rule Inf_greatest)
+ fix x
+ assume "x \<in> ?H"
+ then have "?a \<le> x" by (rule Inf_lower)
+ with `mono f` have "f ?a \<le> f x" ..
+ also from `x \<in> ?H` have "\<dots> \<le> x" ..
+ finally show "f ?a \<le> x" .
+ qed
+ show "?a \<le> f ?a"
+ proof (rule Inf_lower)
+ from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
+ then show "f ?a \<in> ?H" ..
+ qed
+ qed
+qed
+
+end
--- a/src/HOL/Isar_examples/MutilatedCheckerboard.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,299 +0,0 @@
-(* Title: HOL/Isar_examples/MutilatedCheckerboard.thy
- ID: $Id$
- Author: Markus Wenzel, TU Muenchen (Isar document)
- Lawrence C Paulson, Cambridge University Computer Laboratory (original scripts)
-*)
-
-header {* The Mutilated Checker Board Problem *}
-
-theory MutilatedCheckerboard imports Main begin
-
-text {*
- The Mutilated Checker Board Problem, formalized inductively. See
- \cite{paulson-mutilated-board} and
- \url{http://isabelle.in.tum.de/library/HOL/Induct/Mutil.html} for the
- original tactic script version.
-*}
-
-subsection {* Tilings *}
-
-inductive_set
- tiling :: "'a set set => 'a set set"
- for A :: "'a set set"
- where
- empty: "{} : tiling A"
- | Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
-
-
-text "The union of two disjoint tilings is a tiling."
-
-lemma tiling_Un:
- assumes "t : tiling A" and "u : tiling A" and "t Int u = {}"
- shows "t Un u : tiling A"
-proof -
- let ?T = "tiling A"
- from `t : ?T` and `t Int u = {}`
- show "t Un u : ?T"
- proof (induct t)
- case empty
- with `u : ?T` show "{} Un u : ?T" by simp
- next
- case (Un a t)
- show "(a Un t) Un u : ?T"
- proof -
- have "a Un (t Un u) : ?T"
- using `a : A`
- proof (rule tiling.Un)
- from `(a Un t) Int u = {}` have "t Int u = {}" by blast
- then show "t Un u: ?T" by (rule Un)
- from `a <= - t` and `(a Un t) Int u = {}`
- show "a <= - (t Un u)" by blast
- qed
- also have "a Un (t Un u) = (a Un t) Un u"
- by (simp only: Un_assoc)
- finally show ?thesis .
- qed
- qed
-qed
-
-
-subsection {* Basic properties of ``below'' *}
-
-constdefs
- below :: "nat => nat set"
- "below n == {i. i < n}"
-
-lemma below_less_iff [iff]: "(i: below k) = (i < k)"
- by (simp add: below_def)
-
-lemma below_0: "below 0 = {}"
- by (simp add: below_def)
-
-lemma Sigma_Suc1:
- "m = n + 1 ==> below m <*> B = ({n} <*> B) Un (below n <*> B)"
- by (simp add: below_def less_Suc_eq) blast
-
-lemma Sigma_Suc2:
- "m = n + 2 ==> A <*> below m =
- (A <*> {n}) Un (A <*> {n + 1}) Un (A <*> below n)"
- by (auto simp add: below_def)
-
-lemmas Sigma_Suc = Sigma_Suc1 Sigma_Suc2
-
-
-subsection {* Basic properties of ``evnodd'' *}
-
-constdefs
- evnodd :: "(nat * nat) set => nat => (nat * nat) set"
- "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}"
-
-lemma evnodd_iff:
- "(i, j): evnodd A b = ((i, j): A & (i + j) mod 2 = b)"
- by (simp add: evnodd_def)
-
-lemma evnodd_subset: "evnodd A b <= A"
- by (unfold evnodd_def, rule Int_lower1)
-
-lemma evnoddD: "x : evnodd A b ==> x : A"
- by (rule subsetD, rule evnodd_subset)
-
-lemma evnodd_finite: "finite A ==> finite (evnodd A b)"
- by (rule finite_subset, rule evnodd_subset)
-
-lemma evnodd_Un: "evnodd (A Un B) b = evnodd A b Un evnodd B b"
- by (unfold evnodd_def) blast
-
-lemma evnodd_Diff: "evnodd (A - B) b = evnodd A b - evnodd B b"
- by (unfold evnodd_def) blast
-
-lemma evnodd_empty: "evnodd {} b = {}"
- by (simp add: evnodd_def)
-
-lemma evnodd_insert: "evnodd (insert (i, j) C) b =
- (if (i + j) mod 2 = b
- then insert (i, j) (evnodd C b) else evnodd C b)"
- by (simp add: evnodd_def) blast
-
-
-subsection {* Dominoes *}
-
-inductive_set
- domino :: "(nat * nat) set set"
- where
- horiz: "{(i, j), (i, j + 1)} : domino"
- | vertl: "{(i, j), (i + 1, j)} : domino"
-
-lemma dominoes_tile_row:
- "{i} <*> below (2 * n) : tiling domino"
- (is "?B n : ?T")
-proof (induct n)
- case 0
- show ?case by (simp add: below_0 tiling.empty)
-next
- case (Suc n)
- let ?a = "{i} <*> {2 * n + 1} Un {i} <*> {2 * n}"
- have "?B (Suc n) = ?a Un ?B n"
- by (auto simp add: Sigma_Suc Un_assoc)
- moreover have "... : ?T"
- proof (rule tiling.Un)
- have "{(i, 2 * n), (i, 2 * n + 1)} : domino"
- by (rule domino.horiz)
- also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast
- finally show "... : domino" .
- show "?B n : ?T" by (rule Suc)
- show "?a <= - ?B n" by blast
- qed
- ultimately show ?case by simp
-qed
-
-lemma dominoes_tile_matrix:
- "below m <*> below (2 * n) : tiling domino"
- (is "?B m : ?T")
-proof (induct m)
- case 0
- show ?case by (simp add: below_0 tiling.empty)
-next
- case (Suc m)
- let ?t = "{m} <*> below (2 * n)"
- have "?B (Suc m) = ?t Un ?B m" by (simp add: Sigma_Suc)
- moreover have "... : ?T"
- proof (rule tiling_Un)
- show "?t : ?T" by (rule dominoes_tile_row)
- show "?B m : ?T" by (rule Suc)
- show "?t Int ?B m = {}" by blast
- qed
- ultimately show ?case by simp
-qed
-
-lemma domino_singleton:
- assumes d: "d : domino" and "b < 2"
- shows "EX i j. evnodd d b = {(i, j)}" (is "?P d")
- using d
-proof induct
- from `b < 2` have b_cases: "b = 0 | b = 1" by arith
- fix i j
- note [simp] = evnodd_empty evnodd_insert mod_Suc
- from b_cases show "?P {(i, j), (i, j + 1)}" by rule auto
- from b_cases show "?P {(i, j), (i + 1, j)}" by rule auto
-qed
-
-lemma domino_finite:
- assumes d: "d: domino"
- shows "finite d"
- using d
-proof induct
- fix i j :: nat
- show "finite {(i, j), (i, j + 1)}" by (intro finite.intros)
- show "finite {(i, j), (i + 1, j)}" by (intro finite.intros)
-qed
-
-
-subsection {* Tilings of dominoes *}
-
-lemma tiling_domino_finite:
- assumes t: "t : tiling domino" (is "t : ?T")
- shows "finite t" (is "?F t")
- using t
-proof induct
- show "?F {}" by (rule finite.emptyI)
- fix a t assume "?F t"
- assume "a : domino" then have "?F a" by (rule domino_finite)
- from this and `?F t` show "?F (a Un t)" by (rule finite_UnI)
-qed
-
-lemma tiling_domino_01:
- assumes t: "t : tiling domino" (is "t : ?T")
- shows "card (evnodd t 0) = card (evnodd t 1)"
- using t
-proof induct
- case empty
- show ?case by (simp add: evnodd_def)
-next
- case (Un a t)
- let ?e = evnodd
- note hyp = `card (?e t 0) = card (?e t 1)`
- and at = `a <= - t`
- have card_suc:
- "!!b. b < 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
- proof -
- fix b :: nat assume "b < 2"
- have "?e (a Un t) b = ?e a b Un ?e t b" by (rule evnodd_Un)
- also obtain i j where e: "?e a b = {(i, j)}"
- proof -
- from `a \<in> domino` and `b < 2`
- have "EX i j. ?e a b = {(i, j)}" by (rule domino_singleton)
- then show ?thesis by (blast intro: that)
- qed
- moreover have "... Un ?e t b = insert (i, j) (?e t b)" by simp
- moreover have "card ... = Suc (card (?e t b))"
- proof (rule card_insert_disjoint)
- from `t \<in> tiling domino` have "finite t"
- by (rule tiling_domino_finite)
- then show "finite (?e t b)"
- by (rule evnodd_finite)
- from e have "(i, j) : ?e a b" by simp
- with at show "(i, j) ~: ?e t b" by (blast dest: evnoddD)
- qed
- ultimately show "?thesis b" by simp
- qed
- then have "card (?e (a Un t) 0) = Suc (card (?e t 0))" by simp
- also from hyp have "card (?e t 0) = card (?e t 1)" .
- also from card_suc have "Suc ... = card (?e (a Un t) 1)"
- by simp
- finally show ?case .
-qed
-
-
-subsection {* Main theorem *}
-
-constdefs
- mutilated_board :: "nat => nat => (nat * nat) set"
- "mutilated_board m n ==
- below (2 * (m + 1)) <*> below (2 * (n + 1))
- - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
-
-theorem mutil_not_tiling: "mutilated_board m n ~: tiling domino"
-proof (unfold mutilated_board_def)
- let ?T = "tiling domino"
- let ?t = "below (2 * (m + 1)) <*> below (2 * (n + 1))"
- let ?t' = "?t - {(0, 0)}"
- let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}"
-
- show "?t'' ~: ?T"
- proof
- have t: "?t : ?T" by (rule dominoes_tile_matrix)
- assume t'': "?t'' : ?T"
-
- let ?e = evnodd
- have fin: "finite (?e ?t 0)"
- by (rule evnodd_finite, rule tiling_domino_finite, rule t)
-
- note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff
- have "card (?e ?t'' 0) < card (?e ?t' 0)"
- proof -
- have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)})
- < card (?e ?t' 0)"
- proof (rule card_Diff1_less)
- from _ fin show "finite (?e ?t' 0)"
- by (rule finite_subset) auto
- show "(2 * m + 1, 2 * n + 1) : ?e ?t' 0" by simp
- qed
- then show ?thesis by simp
- qed
- also have "... < card (?e ?t 0)"
- proof -
- have "(0, 0) : ?e ?t 0" by simp
- with fin have "card (?e ?t 0 - {(0, 0)}) < card (?e ?t 0)"
- by (rule card_Diff1_less)
- then show ?thesis by simp
- qed
- also from t have "... = card (?e ?t 1)"
- by (rule tiling_domino_01)
- also have "?e ?t 1 = ?e ?t'' 1" by simp
- also from t'' have "card ... = card (?e ?t'' 0)"
- by (rule tiling_domino_01 [symmetric])
- finally have "... < ..." . then show False ..
- qed
-qed
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Mutilated_Checkerboard.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,300 @@
+(* Title: HOL/Isar_examples/Mutilated_Checkerboard.thy
+ Author: Markus Wenzel, TU Muenchen (Isar document)
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory (original scripts)
+*)
+
+header {* The Mutilated Checker Board Problem *}
+
+theory Mutilated_Checkerboard
+imports Main
+begin
+
+text {*
+ The Mutilated Checker Board Problem, formalized inductively. See
+ \cite{paulson-mutilated-board} and
+ \url{http://isabelle.in.tum.de/library/HOL/Induct/Mutil.html} for the
+ original tactic script version.
+*}
+
+subsection {* Tilings *}
+
+inductive_set
+ tiling :: "'a set set => 'a set set"
+ for A :: "'a set set"
+ where
+ empty: "{} : tiling A"
+ | Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
+
+
+text "The union of two disjoint tilings is a tiling."
+
+lemma tiling_Un:
+ assumes "t : tiling A" and "u : tiling A" and "t Int u = {}"
+ shows "t Un u : tiling A"
+proof -
+ let ?T = "tiling A"
+ from `t : ?T` and `t Int u = {}`
+ show "t Un u : ?T"
+ proof (induct t)
+ case empty
+ with `u : ?T` show "{} Un u : ?T" by simp
+ next
+ case (Un a t)
+ show "(a Un t) Un u : ?T"
+ proof -
+ have "a Un (t Un u) : ?T"
+ using `a : A`
+ proof (rule tiling.Un)
+ from `(a Un t) Int u = {}` have "t Int u = {}" by blast
+ then show "t Un u: ?T" by (rule Un)
+ from `a <= - t` and `(a Un t) Int u = {}`
+ show "a <= - (t Un u)" by blast
+ qed
+ also have "a Un (t Un u) = (a Un t) Un u"
+ by (simp only: Un_assoc)
+ finally show ?thesis .
+ qed
+ qed
+qed
+
+
+subsection {* Basic properties of ``below'' *}
+
+constdefs
+ below :: "nat => nat set"
+ "below n == {i. i < n}"
+
+lemma below_less_iff [iff]: "(i: below k) = (i < k)"
+ by (simp add: below_def)
+
+lemma below_0: "below 0 = {}"
+ by (simp add: below_def)
+
+lemma Sigma_Suc1:
+ "m = n + 1 ==> below m <*> B = ({n} <*> B) Un (below n <*> B)"
+ by (simp add: below_def less_Suc_eq) blast
+
+lemma Sigma_Suc2:
+ "m = n + 2 ==> A <*> below m =
+ (A <*> {n}) Un (A <*> {n + 1}) Un (A <*> below n)"
+ by (auto simp add: below_def)
+
+lemmas Sigma_Suc = Sigma_Suc1 Sigma_Suc2
+
+
+subsection {* Basic properties of ``evnodd'' *}
+
+constdefs
+ evnodd :: "(nat * nat) set => nat => (nat * nat) set"
+ "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}"
+
+lemma evnodd_iff:
+ "(i, j): evnodd A b = ((i, j): A & (i + j) mod 2 = b)"
+ by (simp add: evnodd_def)
+
+lemma evnodd_subset: "evnodd A b <= A"
+ by (unfold evnodd_def, rule Int_lower1)
+
+lemma evnoddD: "x : evnodd A b ==> x : A"
+ by (rule subsetD, rule evnodd_subset)
+
+lemma evnodd_finite: "finite A ==> finite (evnodd A b)"
+ by (rule finite_subset, rule evnodd_subset)
+
+lemma evnodd_Un: "evnodd (A Un B) b = evnodd A b Un evnodd B b"
+ by (unfold evnodd_def) blast
+
+lemma evnodd_Diff: "evnodd (A - B) b = evnodd A b - evnodd B b"
+ by (unfold evnodd_def) blast
+
+lemma evnodd_empty: "evnodd {} b = {}"
+ by (simp add: evnodd_def)
+
+lemma evnodd_insert: "evnodd (insert (i, j) C) b =
+ (if (i + j) mod 2 = b
+ then insert (i, j) (evnodd C b) else evnodd C b)"
+ by (simp add: evnodd_def) blast
+
+
+subsection {* Dominoes *}
+
+inductive_set
+ domino :: "(nat * nat) set set"
+ where
+ horiz: "{(i, j), (i, j + 1)} : domino"
+ | vertl: "{(i, j), (i + 1, j)} : domino"
+
+lemma dominoes_tile_row:
+ "{i} <*> below (2 * n) : tiling domino"
+ (is "?B n : ?T")
+proof (induct n)
+ case 0
+ show ?case by (simp add: below_0 tiling.empty)
+next
+ case (Suc n)
+ let ?a = "{i} <*> {2 * n + 1} Un {i} <*> {2 * n}"
+ have "?B (Suc n) = ?a Un ?B n"
+ by (auto simp add: Sigma_Suc Un_assoc)
+ moreover have "... : ?T"
+ proof (rule tiling.Un)
+ have "{(i, 2 * n), (i, 2 * n + 1)} : domino"
+ by (rule domino.horiz)
+ also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast
+ finally show "... : domino" .
+ show "?B n : ?T" by (rule Suc)
+ show "?a <= - ?B n" by blast
+ qed
+ ultimately show ?case by simp
+qed
+
+lemma dominoes_tile_matrix:
+ "below m <*> below (2 * n) : tiling domino"
+ (is "?B m : ?T")
+proof (induct m)
+ case 0
+ show ?case by (simp add: below_0 tiling.empty)
+next
+ case (Suc m)
+ let ?t = "{m} <*> below (2 * n)"
+ have "?B (Suc m) = ?t Un ?B m" by (simp add: Sigma_Suc)
+ moreover have "... : ?T"
+ proof (rule tiling_Un)
+ show "?t : ?T" by (rule dominoes_tile_row)
+ show "?B m : ?T" by (rule Suc)
+ show "?t Int ?B m = {}" by blast
+ qed
+ ultimately show ?case by simp
+qed
+
+lemma domino_singleton:
+ assumes d: "d : domino" and "b < 2"
+ shows "EX i j. evnodd d b = {(i, j)}" (is "?P d")
+ using d
+proof induct
+ from `b < 2` have b_cases: "b = 0 | b = 1" by arith
+ fix i j
+ note [simp] = evnodd_empty evnodd_insert mod_Suc
+ from b_cases show "?P {(i, j), (i, j + 1)}" by rule auto
+ from b_cases show "?P {(i, j), (i + 1, j)}" by rule auto
+qed
+
+lemma domino_finite:
+ assumes d: "d: domino"
+ shows "finite d"
+ using d
+proof induct
+ fix i j :: nat
+ show "finite {(i, j), (i, j + 1)}" by (intro finite.intros)
+ show "finite {(i, j), (i + 1, j)}" by (intro finite.intros)
+qed
+
+
+subsection {* Tilings of dominoes *}
+
+lemma tiling_domino_finite:
+ assumes t: "t : tiling domino" (is "t : ?T")
+ shows "finite t" (is "?F t")
+ using t
+proof induct
+ show "?F {}" by (rule finite.emptyI)
+ fix a t assume "?F t"
+ assume "a : domino" then have "?F a" by (rule domino_finite)
+ from this and `?F t` show "?F (a Un t)" by (rule finite_UnI)
+qed
+
+lemma tiling_domino_01:
+ assumes t: "t : tiling domino" (is "t : ?T")
+ shows "card (evnodd t 0) = card (evnodd t 1)"
+ using t
+proof induct
+ case empty
+ show ?case by (simp add: evnodd_def)
+next
+ case (Un a t)
+ let ?e = evnodd
+ note hyp = `card (?e t 0) = card (?e t 1)`
+ and at = `a <= - t`
+ have card_suc:
+ "!!b. b < 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
+ proof -
+ fix b :: nat assume "b < 2"
+ have "?e (a Un t) b = ?e a b Un ?e t b" by (rule evnodd_Un)
+ also obtain i j where e: "?e a b = {(i, j)}"
+ proof -
+ from `a \<in> domino` and `b < 2`
+ have "EX i j. ?e a b = {(i, j)}" by (rule domino_singleton)
+ then show ?thesis by (blast intro: that)
+ qed
+ moreover have "... Un ?e t b = insert (i, j) (?e t b)" by simp
+ moreover have "card ... = Suc (card (?e t b))"
+ proof (rule card_insert_disjoint)
+ from `t \<in> tiling domino` have "finite t"
+ by (rule tiling_domino_finite)
+ then show "finite (?e t b)"
+ by (rule evnodd_finite)
+ from e have "(i, j) : ?e a b" by simp
+ with at show "(i, j) ~: ?e t b" by (blast dest: evnoddD)
+ qed
+ ultimately show "?thesis b" by simp
+ qed
+ then have "card (?e (a Un t) 0) = Suc (card (?e t 0))" by simp
+ also from hyp have "card (?e t 0) = card (?e t 1)" .
+ also from card_suc have "Suc ... = card (?e (a Un t) 1)"
+ by simp
+ finally show ?case .
+qed
+
+
+subsection {* Main theorem *}
+
+constdefs
+ mutilated_board :: "nat => nat => (nat * nat) set"
+ "mutilated_board m n ==
+ below (2 * (m + 1)) <*> below (2 * (n + 1))
+ - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
+
+theorem mutil_not_tiling: "mutilated_board m n ~: tiling domino"
+proof (unfold mutilated_board_def)
+ let ?T = "tiling domino"
+ let ?t = "below (2 * (m + 1)) <*> below (2 * (n + 1))"
+ let ?t' = "?t - {(0, 0)}"
+ let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}"
+
+ show "?t'' ~: ?T"
+ proof
+ have t: "?t : ?T" by (rule dominoes_tile_matrix)
+ assume t'': "?t'' : ?T"
+
+ let ?e = evnodd
+ have fin: "finite (?e ?t 0)"
+ by (rule evnodd_finite, rule tiling_domino_finite, rule t)
+
+ note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff
+ have "card (?e ?t'' 0) < card (?e ?t' 0)"
+ proof -
+ have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)})
+ < card (?e ?t' 0)"
+ proof (rule card_Diff1_less)
+ from _ fin show "finite (?e ?t' 0)"
+ by (rule finite_subset) auto
+ show "(2 * m + 1, 2 * n + 1) : ?e ?t' 0" by simp
+ qed
+ then show ?thesis by simp
+ qed
+ also have "... < card (?e ?t 0)"
+ proof -
+ have "(0, 0) : ?e ?t 0" by simp
+ with fin have "card (?e ?t 0 - {(0, 0)}) < card (?e ?t 0)"
+ by (rule card_Diff1_less)
+ then show ?thesis by simp
+ qed
+ also from t have "... = card (?e ?t 1)"
+ by (rule tiling_domino_01)
+ also have "?e ?t 1 = ?e ?t'' 1" by simp
+ also from t'' have "card ... = card (?e ?t'' 0)"
+ by (rule tiling_domino_01 [symmetric])
+ finally have "... < ..." . then show False ..
+ qed
+qed
+
+end
--- a/src/HOL/Isar_examples/NestedDatatype.thy Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,87 +0,0 @@
-
-(* $Id$ *)
-
-header {* Nested datatypes *}
-
-theory NestedDatatype imports Main begin
-
-subsection {* Terms and substitution *}
-
-datatype ('a, 'b) "term" =
- Var 'a
- | App 'b "('a, 'b) term list"
-
-consts
- subst_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term"
- subst_term_list ::
- "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list"
-
-primrec (subst)
- "subst_term f (Var a) = f a"
- "subst_term f (App b ts) = App b (subst_term_list f ts)"
- "subst_term_list f [] = []"
- "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
-
-
-text {*
- \medskip A simple lemma about composition of substitutions.
-*}
-
-lemma "subst_term (subst_term f1 o f2) t =
- subst_term f1 (subst_term f2 t)"
- and "subst_term_list (subst_term f1 o f2) ts =
- subst_term_list f1 (subst_term_list f2 ts)"
- by (induct t and ts) simp_all
-
-lemma "subst_term (subst_term f1 o f2) t =
- subst_term f1 (subst_term f2 t)"
-proof -
- let "?P t" = ?thesis
- let ?Q = "\<lambda>ts. subst_term_list (subst_term f1 o f2) ts =
- subst_term_list f1 (subst_term_list f2 ts)"
- show ?thesis
- proof (induct t)
- fix a show "?P (Var a)" by simp
- next
- fix b ts assume "?Q ts"
- then show "?P (App b ts)"
- by (simp only: subst.simps)
- next
- show "?Q []" by simp
- next
- fix t ts
- assume "?P t" "?Q ts" then show "?Q (t # ts)"
- by (simp only: subst.simps)
- qed
-qed
-
-
-subsection {* Alternative induction *}
-
-theorem term_induct' [case_names Var App]:
- assumes var: "!!a. P (Var a)"
- and app: "!!b ts. list_all P ts ==> P (App b ts)"
- shows "P t"
-proof (induct t)
- fix a show "P (Var a)" by (rule var)
-next
- fix b t ts assume "list_all P ts"
- then show "P (App b ts)" by (rule app)
-next
- show "list_all P []" by simp
-next
- fix t ts assume "P t" "list_all P ts"
- then show "list_all P (t # ts)" by simp
-qed
-
-lemma
- "subst_term (subst_term f1 o f2) t = subst_term f1 (subst_term f2 t)"
-proof (induct t rule: term_induct')
- case (Var a)
- show ?case by (simp add: o_def)
-next
- case (App b ts)
- then show ?case by (induct ts) simp_all
-qed
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_examples/Nested_Datatype.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,86 @@
+header {* Nested datatypes *}
+
+theory Nested_Datatype
+imports Main
+begin
+
+subsection {* Terms and substitution *}
+
+datatype ('a, 'b) "term" =
+ Var 'a
+ | App 'b "('a, 'b) term list"
+
+consts
+ subst_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term"
+ subst_term_list ::
+ "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list"
+
+primrec (subst)
+ "subst_term f (Var a) = f a"
+ "subst_term f (App b ts) = App b (subst_term_list f ts)"
+ "subst_term_list f [] = []"
+ "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
+
+
+text {*
+ \medskip A simple lemma about composition of substitutions.
+*}
+
+lemma "subst_term (subst_term f1 o f2) t =
+ subst_term f1 (subst_term f2 t)"
+ and "subst_term_list (subst_term f1 o f2) ts =
+ subst_term_list f1 (subst_term_list f2 ts)"
+ by (induct t and ts) simp_all
+
+lemma "subst_term (subst_term f1 o f2) t =
+ subst_term f1 (subst_term f2 t)"
+proof -
+ let "?P t" = ?thesis
+ let ?Q = "\<lambda>ts. subst_term_list (subst_term f1 o f2) ts =
+ subst_term_list f1 (subst_term_list f2 ts)"
+ show ?thesis
+ proof (induct t)
+ fix a show "?P (Var a)" by simp
+ next
+ fix b ts assume "?Q ts"
+ then show "?P (App b ts)"
+ by (simp only: subst.simps)
+ next
+ show "?Q []" by simp
+ next
+ fix t ts
+ assume "?P t" "?Q ts" then show "?Q (t # ts)"
+ by (simp only: subst.simps)
+ qed
+qed
+
+
+subsection {* Alternative induction *}
+
+theorem term_induct' [case_names Var App]:
+ assumes var: "!!a. P (Var a)"
+ and app: "!!b ts. list_all P ts ==> P (App b ts)"
+ shows "P t"
+proof (induct t)
+ fix a show "P (Var a)" by (rule var)
+next
+ fix b t ts assume "list_all P ts"
+ then show "P (App b ts)" by (rule app)
+next
+ show "list_all P []" by simp
+next
+ fix t ts assume "P t" "list_all P ts"
+ then show "list_all P (t # ts)" by simp
+qed
+
+lemma
+ "subst_term (subst_term f1 o f2) t = subst_term f1 (subst_term f2 t)"
+proof (induct t rule: term_induct')
+ case (Var a)
+ show ?case by (simp add: o_def)
+next
+ case (App b ts)
+ then show ?case by (induct ts) simp_all
+qed
+
+end
--- a/src/HOL/Isar_examples/Peirce.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Peirce.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,11 +1,12 @@
(* Title: HOL/Isar_examples/Peirce.thy
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
*)
header {* Peirce's Law *}
-theory Peirce imports Main begin
+theory Peirce
+imports Main
+begin
text {*
We consider Peirce's Law: $((A \impl B) \impl A) \impl A$. This is
--- a/src/HOL/Isar_examples/Puzzle.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Puzzle.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,7 +1,8 @@
-
header {* An old chestnut *}
-theory Puzzle imports Main begin
+theory Puzzle
+imports Main
+begin
text_raw {*
\footnote{A question from ``Bundeswettbewerb Mathematik''. Original
--- a/src/HOL/Isar_examples/ROOT.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/ROOT.ML Tue Jun 23 15:48:55 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Isar_examples/ROOT.ML
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
@@ -8,16 +7,16 @@
no_document use_thys ["../NumberTheory/Primes", "../NumberTheory/Fibonacci"];
use_thys [
- "BasicLogic",
+ "Basic_Logic",
"Cantor",
"Peirce",
"Drinker",
- "ExprCompiler",
+ "Expr_Compiler",
"Group",
"Summation",
- "KnasterTarski",
- "MutilatedCheckerboard",
+ "Knaster_Tarski",
+ "Mutilated_Checkerboard",
"Puzzle",
- "NestedDatatype",
- "HoareEx"
+ "Nested_Datatype",
+ "Hoare_Ex"
];
--- a/src/HOL/Isar_examples/Summation.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/Summation.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Isar_examples/Summation.thy
- ID: $Id$
Author: Markus Wenzel
*)
--- a/src/HOL/Isar_examples/document/root.tex Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Isar_examples/document/root.tex Tue Jun 23 15:48:55 2009 +0100
@@ -1,6 +1,3 @@
-
-% $Id$
-
\input{style}
\hyphenation{Isabelle}
--- a/src/HOL/Library/Abstract_Rat.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Library/Abstract_Rat.thy Tue Jun 23 15:48:55 2009 +0100
@@ -21,17 +21,17 @@
definition
isnormNum :: "Num \<Rightarrow> bool"
where
- "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> zgcd a b = 1))"
+ "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"
definition
normNum :: "Num \<Rightarrow> Num"
where
"normNum = (\<lambda>(a,b). (if a=0 \<or> b = 0 then (0,0) else
- (let g = zgcd a b
+ (let g = gcd a b
in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
-declare zgcd_zdvd1[presburger]
-declare zgcd_zdvd2[presburger]
+declare int_gcd_dvd1[presburger]
+declare int_gcd_dvd2[presburger]
lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
proof -
have " \<exists> a b. x = (a,b)" by auto
@@ -39,19 +39,19 @@
{assume "a=0 \<or> b = 0" hence ?thesis by (simp add: normNum_def isnormNum_def)}
moreover
{assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
- let ?g = "zgcd a b"
+ let ?g = "gcd a b"
let ?a' = "a div ?g"
let ?b' = "b div ?g"
- let ?g' = "zgcd ?a' ?b'"
- from anz bnz have "?g \<noteq> 0" by simp with zgcd_pos[of a b]
+ let ?g' = "gcd ?a' ?b'"
+ from anz bnz have "?g \<noteq> 0" by simp with int_gcd_ge_0[of a b]
have gpos: "?g > 0" by arith
have gdvd: "?g dvd a" "?g dvd b" by arith+
from zdvd_mult_div_cancel[OF gdvd(1)] zdvd_mult_div_cancel[OF gdvd(2)]
anz bnz
- have nz':"?a' \<noteq> 0" "?b' \<noteq> 0"
- by - (rule notI,simp add:zgcd_def)+
+ have nz':"?a' \<noteq> 0" "?b' \<noteq> 0"
+ by - (rule notI, simp)+
from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
- from div_zgcd_relprime[OF stupid] have gp1: "?g' = 1" .
+ from int_div_gcd_coprime[OF stupid] have gp1: "?g' = 1" .
from bnz have "b < 0 \<or> b > 0" by arith
moreover
{assume b: "b > 0"
@@ -67,7 +67,7 @@
have False using b by arith }
hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])
from anz bnz nz' b b' gp1 have ?thesis
- by (simp add: isnormNum_def normNum_def Let_def split_def fst_conv snd_conv)}
+ by (simp add: isnormNum_def normNum_def Let_def split_def)}
ultimately have ?thesis by blast
}
ultimately show ?thesis by blast
@@ -85,7 +85,7 @@
definition
Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "*\<^sub>N" 60)
where
- "Nmul = (\<lambda>(a,b) (a',b'). let g = zgcd (a*a') (b*b')
+ "Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')
in (a*a' div g, b*b' div g))"
definition
@@ -121,11 +121,11 @@
then obtain a b a' b' where ab: "x = (a,b)" and ab': "y = (a',b')" by blast
{assume "a = 0"
hence ?thesis using xn ab ab'
- by (simp add: zgcd_def isnormNum_def Let_def Nmul_def split_def)}
+ by (simp add: isnormNum_def Let_def Nmul_def split_def)}
moreover
{assume "a' = 0"
hence ?thesis using yn ab ab'
- by (simp add: zgcd_def isnormNum_def Let_def Nmul_def split_def)}
+ by (simp add: isnormNum_def Let_def Nmul_def split_def)}
moreover
{assume a: "a \<noteq>0" and a': "a'\<noteq>0"
hence bp: "b > 0" "b' > 0" using xn yn ab ab' by (simp_all add: isnormNum_def)
@@ -137,11 +137,11 @@
lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
by (simp add: Ninv_def isnormNum_def split_def)
- (cases "fst x = 0", auto simp add: zgcd_commute)
+ (cases "fst x = 0", auto simp add: int_gcd_commute)
lemma isnormNum_int[simp]:
"isnormNum 0\<^sub>N" "isnormNum (1::int)\<^sub>N" "i \<noteq> 0 \<Longrightarrow> isnormNum i\<^sub>N"
- by (simp_all add: isnormNum_def zgcd_def)
+ by (simp_all add: isnormNum_def)
text {* Relations over Num *}
@@ -202,8 +202,8 @@
from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: isnormNum_def)
from prems have eq:"a * b' = a'*b"
by (simp add: INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
- from prems have gcd1: "zgcd a b = 1" "zgcd b a = 1" "zgcd a' b' = 1" "zgcd b' a' = 1"
- by (simp_all add: isnormNum_def add: zgcd_commute)
+ from prems have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
+ by (simp_all add: isnormNum_def add: int_gcd_commute)
from eq have raw_dvd: "a dvd a'*b" "b dvd b'*a" "a' dvd a*b'" "b' dvd b*a'"
apply -
apply algebra
@@ -211,8 +211,8 @@
apply simp
apply algebra
done
- from zdvd_dvd_eq[OF bz zrelprime_dvd_mult[OF gcd1(2) raw_dvd(2)]
- zrelprime_dvd_mult[OF gcd1(4) raw_dvd(4)]]
+ from zdvd_dvd_eq[OF bz int_coprime_dvd_mult[OF gcd1(2) raw_dvd(2)]
+ int_coprime_dvd_mult[OF gcd1(4) raw_dvd(4)]]
have eq1: "b = b'" using pos by arith
with eq have "a = a'" using pos by simp
with eq1 have ?rhs by simp}
@@ -258,7 +258,7 @@
by (simp add: INum_def normNum_def split_def Let_def)}
moreover
{assume a: "a\<noteq>0" and b: "b\<noteq>0"
- let ?g = "zgcd a b"
+ let ?g = "gcd a b"
from a b have g: "?g \<noteq> 0"by simp
from of_int_div[OF g, where ?'a = 'a]
have ?thesis by (auto simp add: INum_def normNum_def split_def Let_def)}
@@ -294,11 +294,11 @@
from z aa' bb' have ?thesis
by (simp add: th Nadd_def normNum_def INum_def split_def)}
moreover {assume z: "a * b' + b * a' \<noteq> 0"
- let ?g = "zgcd (a * b' + b * a') (b*b')"
+ let ?g = "gcd (a * b' + b * a') (b*b')"
have gz: "?g \<noteq> 0" using z by simp
have ?thesis using aa' bb' z gz
- of_int_div[where ?'a = 'a, OF gz zgcd_zdvd1[where i="a * b' + b * a'" and j="b*b'"]] of_int_div[where ?'a = 'a,
- OF gz zgcd_zdvd2[where i="a * b' + b * a'" and j="b*b'"]]
+ of_int_div[where ?'a = 'a, OF gz int_gcd_dvd1[where x="a * b' + b * a'" and y="b*b'"]] of_int_div[where ?'a = 'a,
+ OF gz int_gcd_dvd2[where x="a * b' + b * a'" and y="b*b'"]]
by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib)}
ultimately have ?thesis using aa' bb'
by (simp add: Nadd_def INum_def normNum_def x y Let_def) }
@@ -317,10 +317,10 @@
done }
moreover
{assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
- let ?g="zgcd (a*a') (b*b')"
+ let ?g="gcd (a*a') (b*b')"
have gz: "?g \<noteq> 0" using z by simp
- from z of_int_div[where ?'a = 'a, OF gz zgcd_zdvd1[where i="a*a'" and j="b*b'"]]
- of_int_div[where ?'a = 'a , OF gz zgcd_zdvd2[where i="a*a'" and j="b*b'"]]
+ from z of_int_div[where ?'a = 'a, OF gz int_gcd_dvd1[where x="a*a'" and y="b*b'"]]
+ of_int_div[where ?'a = 'a , OF gz int_gcd_dvd2[where x="a*a'" and y="b*b'"]]
have ?thesis by (simp add: Nmul_def x y Let_def INum_def)}
ultimately show ?thesis by blast
qed
@@ -478,7 +478,7 @@
qed
lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
- by (simp add: Nmul_def split_def Let_def zgcd_commute mult_commute)
+ by (simp add: Nmul_def split_def Let_def int_gcd_commute mult_commute)
lemma Nmul_assoc:
assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
--- a/src/HOL/Library/Formal_Power_Series.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1204,10 +1204,6 @@
apply simp
unfolding One_nat_def[symmetric] natlist_trivial_1
apply simp
- unfolding image_Collect[symmetric]
- unfolding Collect_def mem_def
- apply (rule finite_imageI)
- apply blast
done
qed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Fraction_Field.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,274 @@
+(* Title: Fraction_Field.thy
+ Author: Amine Chaieb, University of Cambridge
+*)
+
+header{* A formalization of the fraction field of any integral domain
+ A generalization of Rational.thy from int to any integral domain *}
+
+theory Fraction_Field
+imports Main (* Equiv_Relations Plain *)
+begin
+
+subsection {* General fractions construction *}
+
+subsubsection {* Construction of the type of fractions *}
+
+definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
+ "fractrel == {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
+
+lemma fractrel_iff [simp]:
+ "(x, y) \<in> fractrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
+ by (simp add: fractrel_def)
+
+lemma refl_fractrel: "refl_on {x. snd x \<noteq> 0} fractrel"
+ by (auto simp add: refl_on_def fractrel_def)
+
+lemma sym_fractrel: "sym fractrel"
+ by (simp add: fractrel_def sym_def)
+
+lemma trans_fractrel: "trans fractrel"
+proof (rule transI, unfold split_paired_all)
+ fix a b a' b' a'' b'' :: 'a
+ assume A: "((a, b), (a', b')) \<in> fractrel"
+ assume B: "((a', b'), (a'', b'')) \<in> fractrel"
+ have "b' * (a * b'') = b'' * (a * b')" by (simp add: mult_ac)
+ also from A have "a * b' = a' * b" by auto
+ also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: mult_ac)
+ also from B have "a' * b'' = a'' * b'" by auto
+ also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: mult_ac)
+ finally have "b' * (a * b'') = b' * (a'' * b)" .
+ moreover from B have "b' \<noteq> 0" by auto
+ ultimately have "a * b'' = a'' * b" by simp
+ with A B show "((a, b), (a'', b'')) \<in> fractrel" by auto
+qed
+
+lemma equiv_fractrel: "equiv {x. snd x \<noteq> 0} fractrel"
+ by (rule equiv.intro [OF refl_fractrel sym_fractrel trans_fractrel])
+
+lemmas UN_fractrel = UN_equiv_class [OF equiv_fractrel]
+lemmas UN_fractrel2 = UN_equiv_class2 [OF equiv_fractrel equiv_fractrel]
+
+lemma equiv_fractrel_iff [iff]:
+ assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
+ shows "fractrel `` {x} = fractrel `` {y} \<longleftrightarrow> (x, y) \<in> fractrel"
+ by (rule eq_equiv_class_iff, rule equiv_fractrel) (auto simp add: assms)
+
+typedef 'a fract = "{(x::'a\<times>'a). snd x \<noteq> (0::'a::idom)} // fractrel"
+proof
+ have "(0::'a, 1::'a) \<in> {x. snd x \<noteq> 0}" by simp
+ then show "fractrel `` {(0::'a, 1)} \<in> {x. snd x \<noteq> 0} // fractrel" by (rule quotientI)
+qed
+
+lemma fractrel_in_fract [simp]: "snd x \<noteq> 0 \<Longrightarrow> fractrel `` {x} \<in> fract"
+ by (simp add: fract_def quotientI)
+
+declare Abs_fract_inject [simp] Abs_fract_inverse [simp]
+
+
+subsubsection {* Representation and basic operations *}
+
+definition
+ Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
+ [code del]: "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
+
+code_datatype Fract
+
+lemma Fract_cases [case_names Fract, cases type: fract]:
+ assumes "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> C"
+ shows C
+ using assms by (cases q) (clarsimp simp add: Fract_def fract_def quotient_def)
+
+lemma Fract_induct [case_names Fract, induct type: fract]:
+ assumes "\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)"
+ shows "P q"
+ using assms by (cases q) simp
+
+lemma eq_fract:
+ shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
+ and "\<And>a. Fract a 0 = Fract 0 1"
+ and "\<And>a c. Fract 0 a = Fract 0 c"
+ by (simp_all add: Fract_def)
+
+instantiation fract :: (idom) "{comm_ring_1, power}"
+begin
+
+definition
+ Zero_fract_def [code, code unfold]: "0 = Fract 0 1"
+
+definition
+ One_fract_def [code, code unfold]: "1 = Fract 1 1"
+
+definition
+ add_fract_def [code del]:
+ "q + r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
+ fractrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
+
+lemma add_fract [simp]:
+ assumes "b \<noteq> (0::'a::idom)" and "d \<noteq> 0"
+ shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
+proof -
+ have "(\<lambda>x y. fractrel``{(fst x * snd y + fst y * snd x, snd x * snd y :: 'a)})
+ respects2 fractrel"
+ apply (rule equiv_fractrel [THEN congruent2_commuteI]) apply (auto simp add: algebra_simps)
+ unfolding mult_assoc[symmetric] .
+ with assms show ?thesis by (simp add: Fract_def add_fract_def UN_fractrel2)
+qed
+
+definition
+ minus_fract_def [code del]:
+ "- q = Abs_fract (\<Union>x \<in> Rep_fract q. fractrel `` {(- fst x, snd x)})"
+
+lemma minus_fract [simp, code]: "- Fract a b = Fract (- a) (b::'a::idom)"
+proof -
+ have "(\<lambda>x. fractrel `` {(- fst x, snd x :: 'a)}) respects fractrel"
+ by (simp add: congruent_def)
+ then show ?thesis by (simp add: Fract_def minus_fract_def UN_fractrel)
+qed
+
+lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
+ by (cases "b = 0") (simp_all add: eq_fract)
+
+definition
+ diff_fract_def [code del]: "q - r = q + - (r::'a fract)"
+
+lemma diff_fract [simp]:
+ assumes "b \<noteq> 0" and "d \<noteq> 0"
+ shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
+ using assms by (simp add: diff_fract_def diff_minus)
+
+definition
+ mult_fract_def [code del]:
+ "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
+ fractrel``{(fst x * fst y, snd x * snd y)})"
+
+lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"
+proof -
+ have "(\<lambda>x y. fractrel `` {(fst x * fst y, snd x * snd y :: 'a)}) respects2 fractrel"
+ apply (rule equiv_fractrel [THEN congruent2_commuteI]) apply (auto simp add: algebra_simps)
+ unfolding mult_assoc[symmetric] .
+ then show ?thesis by (simp add: Fract_def mult_fract_def UN_fractrel2)
+qed
+
+lemma mult_fract_cancel:
+ assumes "c \<noteq> 0"
+ shows "Fract (c * a) (c * b) = Fract a b"
+proof -
+ from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def)
+ then show ?thesis by (simp add: mult_fract [symmetric])
+qed
+
+instance proof
+ fix q r s :: "'a fract" show "(q * r) * s = q * (r * s)"
+ by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+next
+ fix q r :: "'a fract" show "q * r = r * q"
+ by (cases q, cases r) (simp add: eq_fract algebra_simps)
+next
+ fix q :: "'a fract" show "1 * q = q"
+ by (cases q) (simp add: One_fract_def eq_fract)
+next
+ fix q r s :: "'a fract" show "(q + r) + s = q + (r + s)"
+ by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+next
+ fix q r :: "'a fract" show "q + r = r + q"
+ by (cases q, cases r) (simp add: eq_fract algebra_simps)
+next
+ fix q :: "'a fract" show "0 + q = q"
+ by (cases q) (simp add: Zero_fract_def eq_fract)
+next
+ fix q :: "'a fract" show "- q + q = 0"
+ by (cases q) (simp add: Zero_fract_def eq_fract)
+next
+ fix q r :: "'a fract" show "q - r = q + - r"
+ by (cases q, cases r) (simp add: eq_fract)
+next
+ fix q r s :: "'a fract" show "(q + r) * s = q * s + r * s"
+ by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+next
+ show "(0::'a fract) \<noteq> 1" by (simp add: Zero_fract_def One_fract_def eq_fract)
+qed
+
+end
+
+lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1"
+ by (induct k) (simp_all add: Zero_fract_def One_fract_def)
+
+lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
+ by (rule of_nat_fract [symmetric])
+
+lemma fract_collapse [code post]:
+ "Fract 0 k = 0"
+ "Fract 1 1 = 1"
+ "Fract k 0 = 0"
+ by (cases "k = 0")
+ (simp_all add: Zero_fract_def One_fract_def eq_fract Fract_def)
+
+lemma fract_expand [code unfold]:
+ "0 = Fract 0 1"
+ "1 = Fract 1 1"
+ by (simp_all add: fract_collapse)
+
+lemma Fract_cases_nonzero [case_names Fract 0]:
+ assumes Fract: "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> C"
+ assumes 0: "q = 0 \<Longrightarrow> C"
+ shows C
+proof (cases "q = 0")
+ case True then show C using 0 by auto
+next
+ case False
+ then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
+ moreover with False have "0 \<noteq> Fract a b" by simp
+ with `b \<noteq> 0` have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
+ with Fract `q = Fract a b` `b \<noteq> 0` show C by auto
+qed
+
+
+
+subsubsection {* The field of rational numbers *}
+
+context idom
+begin
+subclass ring_no_zero_divisors ..
+thm mult_eq_0_iff
+end
+
+instantiation fract :: (idom) "{field, division_by_zero}"
+begin
+
+definition
+ inverse_fract_def [code del]:
+ "inverse q = Abs_fract (\<Union>x \<in> Rep_fract q.
+ fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
+
+
+lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"
+proof -
+ have stupid: "\<And>x. (0::'a) = x \<longleftrightarrow> x = 0" by auto
+ have "(\<lambda>x. fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x :: 'a)}) respects fractrel"
+ by (auto simp add: congruent_def stupid algebra_simps)
+ then show ?thesis by (simp add: Fract_def inverse_fract_def UN_fractrel)
+qed
+
+definition
+ divide_fract_def [code del]: "q / r = q * inverse (r:: 'a fract)"
+
+lemma divide_fract [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
+ by (simp add: divide_fract_def)
+
+instance proof
+ show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
+ (simp add: fract_collapse)
+next
+ fix q :: "'a fract"
+ assume "q \<noteq> 0"
+ then show "inverse q * q = 1" apply (cases q rule: Fract_cases_nonzero)
+ by (simp_all add: mult_fract inverse_fract fract_expand eq_fract mult_commute)
+next
+ fix q r :: "'a fract"
+ show "q / r = q * inverse r" by (simp add: divide_fract_def)
+qed
+
+end
+
+
+end
\ No newline at end of file
--- a/src/HOL/Library/FuncSet.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Library/FuncSet.thy Tue Jun 23 15:48:55 2009 +0100
@@ -51,22 +51,32 @@
subsection{*Basic Properties of @{term Pi}*}
-lemma Pi_I: "(!!x. x \<in> A ==> f x \<in> B x) ==> f \<in> Pi A B"
+lemma Pi_I[intro!]: "(!!x. x \<in> A ==> f x \<in> B x) ==> f \<in> Pi A B"
by (simp add: Pi_def)
+lemma Pi_I'[simp]: "(!!x. x : A --> f x : B x) ==> f : Pi A B"
+by(simp add:Pi_def)
+
lemma funcsetI: "(!!x. x \<in> A ==> f x \<in> B) ==> f \<in> A -> B"
by (simp add: Pi_def)
lemma Pi_mem: "[|f: Pi A B; x \<in> A|] ==> f x \<in> B x"
by (simp add: Pi_def)
+lemma PiE [elim]:
+ "f : Pi A B ==> (f x : B x ==> Q) ==> (x ~: A ==> Q) ==> Q"
+by(auto simp: Pi_def)
+
+lemma funcset_id [simp]: "(\<lambda>x. x) \<in> A \<rightarrow> A"
+ by (auto intro: Pi_I)
+
lemma funcset_mem: "[|f \<in> A -> B; x \<in> A|] ==> f x \<in> B"
by (simp add: Pi_def)
lemma funcset_image: "f \<in> A\<rightarrow>B ==> f ` A \<subseteq> B"
- by (auto simp add: Pi_def)
+by auto
-lemma Pi_eq_empty: "((PI x: A. B x) = {}) = (\<exists>x\<in>A. B(x) = {})"
+lemma Pi_eq_empty[simp]: "((PI x: A. B x) = {}) = (\<exists>x\<in>A. B(x) = {})"
apply (simp add: Pi_def, auto)
txt{*Converse direction requires Axiom of Choice to exhibit a function
picking an element from each non-empty @{term "B x"}*}
@@ -75,33 +85,36 @@
done
lemma Pi_empty [simp]: "Pi {} B = UNIV"
- by (simp add: Pi_def)
+by (simp add: Pi_def)
lemma Pi_UNIV [simp]: "A -> UNIV = UNIV"
+by (simp add: Pi_def)
+(*
+lemma funcset_id [simp]: "(%x. x): A -> A"
by (simp add: Pi_def)
-
+*)
text{*Covariance of Pi-sets in their second argument*}
lemma Pi_mono: "(!!x. x \<in> A ==> B x <= C x) ==> Pi A B <= Pi A C"
- by (simp add: Pi_def, blast)
+by auto
text{*Contravariance of Pi-sets in their first argument*}
lemma Pi_anti_mono: "A' <= A ==> Pi A B <= Pi A' B"
- by (simp add: Pi_def, blast)
+by auto
subsection{*Composition With a Restricted Domain: @{term compose}*}
lemma funcset_compose:
- "[| f \<in> A -> B; g \<in> B -> C |]==> compose A g f \<in> A -> C"
- by (simp add: Pi_def compose_def restrict_def)
+ "[| f \<in> A -> B; g \<in> B -> C |]==> compose A g f \<in> A -> C"
+by (simp add: Pi_def compose_def restrict_def)
lemma compose_assoc:
"[| f \<in> A -> B; g \<in> B -> C; h \<in> C -> D |]
==> compose A h (compose A g f) = compose A (compose B h g) f"
- by (simp add: expand_fun_eq Pi_def compose_def restrict_def)
+by (simp add: expand_fun_eq Pi_def compose_def restrict_def)
lemma compose_eq: "x \<in> A ==> compose A g f x = g(f(x))"
- by (simp add: compose_def restrict_def)
+by (simp add: compose_def restrict_def)
lemma surj_compose: "[| f ` A = B; g ` B = C |] ==> compose A g f ` A = C"
by (auto simp add: image_def compose_eq)
@@ -112,7 +125,7 @@
lemma restrict_in_funcset: "(!!x. x \<in> A ==> f x \<in> B) ==> (\<lambda>x\<in>A. f x) \<in> A -> B"
by (simp add: Pi_def restrict_def)
-lemma restrictI: "(!!x. x \<in> A ==> f x \<in> B x) ==> (\<lambda>x\<in>A. f x) \<in> Pi A B"
+lemma restrictI[intro!]: "(!!x. x \<in> A ==> f x \<in> B x) ==> (\<lambda>x\<in>A. f x) \<in> Pi A B"
by (simp add: Pi_def restrict_def)
lemma restrict_apply [simp]:
@@ -121,7 +134,7 @@
lemma restrict_ext:
"(!!x. x \<in> A ==> f x = g x) ==> (\<lambda>x\<in>A. f x) = (\<lambda>x\<in>A. g x)"
- by (simp add: expand_fun_eq Pi_def Pi_def restrict_def)
+ by (simp add: expand_fun_eq Pi_def restrict_def)
lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A"
by (simp add: inj_on_def restrict_def)
@@ -144,68 +157,66 @@
the theorems belong here, or need at least @{term Hilbert_Choice}.*}
lemma bij_betw_imp_funcset: "bij_betw f A B \<Longrightarrow> f \<in> A \<rightarrow> B"
- by (auto simp add: bij_betw_def inj_on_Inv Pi_def)
+by (auto simp add: bij_betw_def inj_on_Inv)
lemma inj_on_compose:
- "[| bij_betw f A B; inj_on g B |] ==> inj_on (compose A g f) A"
- by (auto simp add: bij_betw_def inj_on_def compose_eq)
+ "[| bij_betw f A B; inj_on g B |] ==> inj_on (compose A g f) A"
+by (auto simp add: bij_betw_def inj_on_def compose_eq)
lemma bij_betw_compose:
- "[| bij_betw f A B; bij_betw g B C |] ==> bij_betw (compose A g f) A C"
- apply (simp add: bij_betw_def compose_eq inj_on_compose)
- apply (auto simp add: compose_def image_def)
- done
+ "[| bij_betw f A B; bij_betw g B C |] ==> bij_betw (compose A g f) A C"
+apply (simp add: bij_betw_def compose_eq inj_on_compose)
+apply (auto simp add: compose_def image_def)
+done
lemma bij_betw_restrict_eq [simp]:
- "bij_betw (restrict f A) A B = bij_betw f A B"
- by (simp add: bij_betw_def)
+ "bij_betw (restrict f A) A B = bij_betw f A B"
+by (simp add: bij_betw_def)
subsection{*Extensionality*}
lemma extensional_arb: "[|f \<in> extensional A; x\<notin> A|] ==> f x = undefined"
- by (simp add: extensional_def)
+by (simp add: extensional_def)
lemma restrict_extensional [simp]: "restrict f A \<in> extensional A"
- by (simp add: restrict_def extensional_def)
+by (simp add: restrict_def extensional_def)
lemma compose_extensional [simp]: "compose A f g \<in> extensional A"
- by (simp add: compose_def)
+by (simp add: compose_def)
lemma extensionalityI:
- "[| f \<in> extensional A; g \<in> extensional A;
+ "[| f \<in> extensional A; g \<in> extensional A;
!!x. x\<in>A ==> f x = g x |] ==> f = g"
- by (force simp add: expand_fun_eq extensional_def)
+by (force simp add: expand_fun_eq extensional_def)
lemma Inv_funcset: "f ` A = B ==> (\<lambda>x\<in>B. Inv A f x) : B -> A"
- by (unfold Inv_def) (fast intro: restrict_in_funcset someI2)
+by (unfold Inv_def) (fast intro: someI2)
lemma compose_Inv_id:
- "bij_betw f A B ==> compose A (\<lambda>y\<in>B. Inv A f y) f = (\<lambda>x\<in>A. x)"
- apply (simp add: bij_betw_def compose_def)
- apply (rule restrict_ext, auto)
- apply (erule subst)
- apply (simp add: Inv_f_f)
- done
+ "bij_betw f A B ==> compose A (\<lambda>y\<in>B. Inv A f y) f = (\<lambda>x\<in>A. x)"
+apply (simp add: bij_betw_def compose_def)
+apply (rule restrict_ext, auto)
+apply (erule subst)
+apply (simp add: Inv_f_f)
+done
lemma compose_id_Inv:
- "f ` A = B ==> compose B f (\<lambda>y\<in>B. Inv A f y) = (\<lambda>x\<in>B. x)"
- apply (simp add: compose_def)
- apply (rule restrict_ext)
- apply (simp add: f_Inv_f)
- done
+ "f ` A = B ==> compose B f (\<lambda>y\<in>B. Inv A f y) = (\<lambda>x\<in>B. x)"
+apply (simp add: compose_def)
+apply (rule restrict_ext)
+apply (simp add: f_Inv_f)
+done
subsection{*Cardinality*}
lemma card_inj: "[|f \<in> A\<rightarrow>B; inj_on f A; finite B|] ==> card(A) \<le> card(B)"
- apply (rule card_inj_on_le)
- apply (auto simp add: Pi_def)
- done
+by (rule card_inj_on_le) auto
lemma card_bij:
- "[|f \<in> A\<rightarrow>B; inj_on f A;
- g \<in> B\<rightarrow>A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)"
- by (blast intro: card_inj order_antisym)
+ "[|f \<in> A\<rightarrow>B; inj_on f A;
+ g \<in> B\<rightarrow>A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)"
+by (blast intro: card_inj order_antisym)
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Legacy_GCD.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,787 @@
+(* Title: HOL/GCD.thy
+ Author: Christophe Tabacznyj and Lawrence C Paulson
+ Copyright 1996 University of Cambridge
+*)
+
+header {* The Greatest Common Divisor *}
+
+theory Legacy_GCD
+imports Main
+begin
+
+text {*
+ See \cite{davenport92}. \bigskip
+*}
+
+subsection {* Specification of GCD on nats *}
+
+definition
+ is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
+ [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
+ (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
+
+text {* Uniqueness *}
+
+lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
+ by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
+
+text {* Connection to divides relation *}
+
+lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
+ by (auto simp add: is_gcd_def)
+
+text {* Commutativity *}
+
+lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
+ by (auto simp add: is_gcd_def)
+
+
+subsection {* GCD on nat by Euclid's algorithm *}
+
+fun
+ gcd :: "nat => nat => nat"
+where
+ "gcd m n = (if n = 0 then m else gcd n (m mod n))"
+lemma gcd_induct [case_names "0" rec]:
+ fixes m n :: nat
+ assumes "\<And>m. P m 0"
+ and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
+ shows "P m n"
+proof (induct m n rule: gcd.induct)
+ case (1 m n) with assms show ?case by (cases "n = 0") simp_all
+qed
+
+lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
+ by simp
+
+lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
+ by simp
+
+lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
+ by simp
+
+lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
+ by simp
+
+lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
+ unfolding One_nat_def by (rule gcd_1)
+
+declare gcd.simps [simp del]
+
+text {*
+ \medskip @{term "gcd m n"} divides @{text m} and @{text n}. The
+ conjunctions don't seem provable separately.
+*}
+
+lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
+ and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
+ apply (induct m n rule: gcd_induct)
+ apply (simp_all add: gcd_non_0)
+ apply (blast dest: dvd_mod_imp_dvd)
+ done
+
+text {*
+ \medskip Maximality: for all @{term m}, @{term n}, @{term k}
+ naturals, if @{term k} divides @{term m} and @{term k} divides
+ @{term n} then @{term k} divides @{term "gcd m n"}.
+*}
+
+lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
+ by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
+
+text {*
+ \medskip Function gcd yields the Greatest Common Divisor.
+*}
+
+lemma is_gcd: "is_gcd m n (gcd m n) "
+ by (simp add: is_gcd_def gcd_greatest)
+
+
+subsection {* Derived laws for GCD *}
+
+lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
+ by (blast intro!: gcd_greatest intro: dvd_trans)
+
+lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
+ by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
+
+lemma gcd_commute: "gcd m n = gcd n m"
+ apply (rule is_gcd_unique)
+ apply (rule is_gcd)
+ apply (subst is_gcd_commute)
+ apply (simp add: is_gcd)
+ done
+
+lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
+ apply (rule is_gcd_unique)
+ apply (rule is_gcd)
+ apply (simp add: is_gcd_def)
+ apply (blast intro: dvd_trans)
+ done
+
+lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
+ by (simp add: gcd_commute)
+
+lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
+ unfolding One_nat_def by (rule gcd_1_left)
+
+text {*
+ \medskip Multiplication laws
+*}
+
+lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
+ -- {* \cite[page 27]{davenport92} *}
+ apply (induct m n rule: gcd_induct)
+ apply simp
+ apply (case_tac "k = 0")
+ apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
+ done
+
+lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
+ apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
+ done
+
+lemma gcd_self [simp, algebra]: "gcd k k = k"
+ apply (rule gcd_mult [of k 1, simplified])
+ done
+
+lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
+ apply (insert gcd_mult_distrib2 [of m k n])
+ apply simp
+ apply (erule_tac t = m in ssubst)
+ apply simp
+ done
+
+lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
+ by (auto intro: relprime_dvd_mult dvd_mult2)
+
+lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
+ apply (rule dvd_anti_sym)
+ apply (rule gcd_greatest)
+ apply (rule_tac n = k in relprime_dvd_mult)
+ apply (simp add: gcd_assoc)
+ apply (simp add: gcd_commute)
+ apply (simp_all add: mult_commute)
+ done
+
+
+text {* \medskip Addition laws *}
+
+lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
+ by (cases "n = 0") (auto simp add: gcd_non_0)
+
+lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
+proof -
+ have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
+ also have "... = gcd (n + m) m" by (simp add: add_commute)
+ also have "... = gcd n m" by simp
+ also have "... = gcd m n" by (rule gcd_commute)
+ finally show ?thesis .
+qed
+
+lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
+ apply (subst add_commute)
+ apply (rule gcd_add2)
+ done
+
+lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
+ by (induct k) (simp_all add: add_assoc)
+
+lemma gcd_dvd_prod: "gcd m n dvd m * n"
+ using mult_dvd_mono [of 1] by auto
+
+text {*
+ \medskip Division by gcd yields rrelatively primes.
+*}
+
+lemma div_gcd_relprime:
+ assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
+ shows "gcd (a div gcd a b) (b div gcd a b) = 1"
+proof -
+ let ?g = "gcd a b"
+ let ?a' = "a div ?g"
+ let ?b' = "b div ?g"
+ let ?g' = "gcd ?a' ?b'"
+ have dvdg: "?g dvd a" "?g dvd b" by simp_all
+ have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
+ from dvdg dvdg' obtain ka kb ka' kb' where
+ kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
+ unfolding dvd_def by blast
+ then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
+ then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+ by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
+ dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
+ have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
+ then have gp: "?g > 0" by simp
+ from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+ with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
+qed
+
+
+lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
+proof(auto)
+ assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
+ from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b]
+ have th: "gcd a b dvd d" by blast
+ from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast
+qed
+
+lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
+ shows "gcd x y = gcd u v"
+proof-
+ from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
+ with gcd_unique[of "gcd u v" x y] show ?thesis by auto
+qed
+
+lemma ind_euclid:
+ assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
+ and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"
+ shows "P a b"
+proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
+ fix n a b
+ assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
+ have "a = b \<or> a < b \<or> b < a" by arith
+ moreover {assume eq: "a= b"
+ from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
+ moreover
+ {assume lt: "a < b"
+ hence "a + b - a < n \<or> a = 0" using H(2) by arith
+ moreover
+ {assume "a =0" with z c have "P a b" by blast }
+ moreover
+ {assume ab: "a + b - a < n"
+ have th0: "a + b - a = a + (b - a)" using lt by arith
+ from add[rule_format, OF H(1)[rule_format, OF ab th0]]
+ have "P a b" by (simp add: th0[symmetric])}
+ ultimately have "P a b" by blast}
+ moreover
+ {assume lt: "a > b"
+ hence "b + a - b < n \<or> b = 0" using H(2) by arith
+ moreover
+ {assume "b =0" with z c have "P a b" by blast }
+ moreover
+ {assume ab: "b + a - b < n"
+ have th0: "b + a - b = b + (a - b)" using lt by arith
+ from add[rule_format, OF H(1)[rule_format, OF ab th0]]
+ have "P b a" by (simp add: th0[symmetric])
+ hence "P a b" using c by blast }
+ ultimately have "P a b" by blast}
+ultimately show "P a b" by blast
+qed
+
+lemma bezout_lemma:
+ assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
+ shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
+using ex
+apply clarsimp
+apply (rule_tac x="d" in exI, simp add: dvd_add)
+apply (case_tac "a * x = b * y + d" , simp_all)
+apply (rule_tac x="x + y" in exI)
+apply (rule_tac x="y" in exI)
+apply algebra
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="x + y" in exI)
+apply algebra
+done
+
+lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
+apply(induct a b rule: ind_euclid)
+apply blast
+apply clarify
+apply (rule_tac x="a" in exI, simp add: dvd_add)
+apply clarsimp
+apply (rule_tac x="d" in exI)
+apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
+apply (rule_tac x="x+y" in exI)
+apply (rule_tac x="y" in exI)
+apply algebra
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="x+y" in exI)
+apply algebra
+done
+
+lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
+using bezout_add[of a b]
+apply clarsimp
+apply (rule_tac x="d" in exI, simp)
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="y" in exI)
+apply auto
+done
+
+
+text {* We can get a stronger version with a nonzeroness assumption. *}
+lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
+
+lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
+ shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
+proof-
+ from nz have ap: "a > 0" by simp
+ from bezout_add[of a b]
+ have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
+ moreover
+ {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
+ from H have ?thesis by blast }
+ moreover
+ {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
+ {assume b0: "b = 0" with H have ?thesis by simp}
+ moreover
+ {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
+ from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
+ moreover
+ {assume db: "d=b"
+ from prems have ?thesis apply simp
+ apply (rule exI[where x = b], simp)
+ apply (rule exI[where x = b])
+ by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
+ moreover
+ {assume db: "d < b"
+ {assume "x=0" hence ?thesis using prems by simp }
+ moreover
+ {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
+
+ from db have "d \<le> b - 1" by simp
+ hence "d*b \<le> b*(b - 1)" by simp
+ with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
+ have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
+ from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
+ hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
+ hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
+ by (simp only: diff_add_assoc[OF dble, of d, symmetric])
+ hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
+ by (simp only: diff_mult_distrib2 add_commute mult_ac)
+ hence ?thesis using H(1,2)
+ apply -
+ apply (rule exI[where x=d], simp)
+ apply (rule exI[where x="(b - 1) * y"])
+ by (rule exI[where x="x*(b - 1) - d"], simp)}
+ ultimately have ?thesis by blast}
+ ultimately have ?thesis by blast}
+ ultimately have ?thesis by blast}
+ ultimately show ?thesis by blast
+qed
+
+
+lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
+proof-
+ let ?g = "gcd a b"
+ from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
+ from d(1,2) have "d dvd ?g" by simp
+ then obtain k where k: "?g = d*k" unfolding dvd_def by blast
+ from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast
+ hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k"
+ by (algebra add: diff_mult_distrib)
+ hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g"
+ by (simp add: k mult_assoc)
+ thus ?thesis by blast
+qed
+
+lemma bezout_gcd_strong: assumes a: "a \<noteq> 0"
+ shows "\<exists>x y. a * x = b * y + gcd a b"
+proof-
+ let ?g = "gcd a b"
+ from bezout_add_strong[OF a, of b]
+ obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
+ from d(1,2) have "d dvd ?g" by simp
+ then obtain k where k: "?g = d*k" unfolding dvd_def by blast
+ from d(3) have "a * x * k = (b * y + d) *k " by algebra
+ hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
+ thus ?thesis by blast
+qed
+
+lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
+by(simp add: gcd_mult_distrib2 mult_commute)
+
+lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
+ (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+ let ?g = "gcd a b"
+ {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
+ from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
+ by blast
+ hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
+ hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k"
+ by (simp only: diff_mult_distrib)
+ hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
+ by (simp add: k[symmetric] mult_assoc)
+ hence ?lhs by blast}
+ moreover
+ {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
+ have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
+ using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
+ from nat_dvd_diff[OF dv(1,2)] nat_dvd_diff[OF dv(3,4)] H
+ have ?rhs by auto}
+ ultimately show ?thesis by blast
+qed
+
+lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
+proof-
+ let ?g = "gcd a b"
+ have dv: "?g dvd a*x" "?g dvd b * y"
+ using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
+ from dvd_add[OF dv] H
+ show ?thesis by auto
+qed
+
+lemma gcd_mult': "gcd b (a * b) = b"
+by (simp add: gcd_mult mult_commute[of a b])
+
+lemma gcd_add: "gcd(a + b) b = gcd a b"
+ "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
+apply (simp_all add: gcd_add1)
+by (simp add: gcd_commute gcd_add1)
+
+lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
+proof-
+ {fix a b assume H: "b \<le> (a::nat)"
+ hence th: "a - b + b = a" by arith
+ from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp}
+ note th = this
+{
+ assume ab: "b \<le> a"
+ from th[OF ab] show "gcd (a - b) b = gcd a b" by blast
+next
+ assume ab: "a \<le> b"
+ from th[OF ab] show "gcd a (b - a) = gcd a b"
+ by (simp add: gcd_commute)}
+qed
+
+
+subsection {* LCM defined by GCD *}
+
+
+definition
+ lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+ lcm_def: "lcm m n = m * n div gcd m n"
+
+lemma prod_gcd_lcm:
+ "m * n = gcd m n * lcm m n"
+ unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
+
+lemma lcm_0 [simp]: "lcm m 0 = 0"
+ unfolding lcm_def by simp
+
+lemma lcm_1 [simp]: "lcm m 1 = m"
+ unfolding lcm_def by simp
+
+lemma lcm_0_left [simp]: "lcm 0 n = 0"
+ unfolding lcm_def by simp
+
+lemma lcm_1_left [simp]: "lcm 1 m = m"
+ unfolding lcm_def by simp
+
+lemma dvd_pos:
+ fixes n m :: nat
+ assumes "n > 0" and "m dvd n"
+ shows "m > 0"
+using assms by (cases m) auto
+
+lemma lcm_least:
+ assumes "m dvd k" and "n dvd k"
+ shows "lcm m n dvd k"
+proof (cases k)
+ case 0 then show ?thesis by auto
+next
+ case (Suc _) then have pos_k: "k > 0" by auto
+ from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
+ with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
+ from assms obtain p where k_m: "k = m * p" using dvd_def by blast
+ from assms obtain q where k_n: "k = n * q" using dvd_def by blast
+ from pos_k k_m have pos_p: "p > 0" by auto
+ from pos_k k_n have pos_q: "q > 0" by auto
+ have "k * k * gcd q p = k * gcd (k * q) (k * p)"
+ by (simp add: mult_ac gcd_mult_distrib2)
+ also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
+ by (simp add: k_m [symmetric] k_n [symmetric])
+ also have "\<dots> = k * p * q * gcd m n"
+ by (simp add: mult_ac gcd_mult_distrib2)
+ finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
+ by (simp only: k_m [symmetric] k_n [symmetric])
+ then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
+ by (simp add: mult_ac)
+ with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
+ by simp
+ with prod_gcd_lcm [of m n]
+ have "lcm m n * gcd q p * gcd m n = k * gcd m n"
+ by (simp add: mult_ac)
+ with pos_gcd have "lcm m n * gcd q p = k" by simp
+ then show ?thesis using dvd_def by auto
+qed
+
+lemma lcm_dvd1 [iff]:
+ "m dvd lcm m n"
+proof (cases m)
+ case 0 then show ?thesis by simp
+next
+ case (Suc _)
+ then have mpos: "m > 0" by simp
+ show ?thesis
+ proof (cases n)
+ case 0 then show ?thesis by simp
+ next
+ case (Suc _)
+ then have npos: "n > 0" by simp
+ have "gcd m n dvd n" by simp
+ then obtain k where "n = gcd m n * k" using dvd_def by auto
+ then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
+ also have "\<dots> = m * k" using mpos npos gcd_zero by simp
+ finally show ?thesis by (simp add: lcm_def)
+ qed
+qed
+
+lemma lcm_dvd2 [iff]:
+ "n dvd lcm m n"
+proof (cases n)
+ case 0 then show ?thesis by simp
+next
+ case (Suc _)
+ then have npos: "n > 0" by simp
+ show ?thesis
+ proof (cases m)
+ case 0 then show ?thesis by simp
+ next
+ case (Suc _)
+ then have mpos: "m > 0" by simp
+ have "gcd m n dvd m" by simp
+ then obtain k where "m = gcd m n * k" using dvd_def by auto
+ then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
+ also have "\<dots> = n * k" using mpos npos gcd_zero by simp
+ finally show ?thesis by (simp add: lcm_def)
+ qed
+qed
+
+lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
+ by (simp add: gcd_commute)
+
+lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
+ apply (subgoal_tac "n = m + (n - m)")
+ apply (erule ssubst, rule gcd_add1_eq, simp)
+ done
+
+
+subsection {* GCD and LCM on integers *}
+
+definition
+ zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
+ "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
+
+lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i"
+by (simp add: zgcd_def int_dvd_iff)
+
+lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j"
+by (simp add: zgcd_def int_dvd_iff)
+
+lemma zgcd_pos: "zgcd i j \<ge> 0"
+by (simp add: zgcd_def)
+
+lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
+by (simp add: zgcd_def gcd_zero)
+
+lemma zgcd_commute: "zgcd i j = zgcd j i"
+unfolding zgcd_def by (simp add: gcd_commute)
+
+lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
+unfolding zgcd_def by simp
+
+lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
+unfolding zgcd_def by simp
+
+ (* should be solved by algebra*)
+lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
+ unfolding zgcd_def
+proof -
+ assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
+ then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
+ from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
+ have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
+ unfolding dvd_def
+ by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
+ from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
+ unfolding dvd_def by blast
+ from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
+ then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
+ then show ?thesis
+ apply (subst abs_dvd_iff [symmetric])
+ apply (subst dvd_abs_iff [symmetric])
+ apply (unfold dvd_def)
+ apply (rule_tac x = "int h'" in exI, simp)
+ done
+qed
+
+lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
+
+lemma zgcd_greatest:
+ assumes "k dvd m" and "k dvd n"
+ shows "k dvd zgcd m n"
+proof -
+ let ?k' = "nat \<bar>k\<bar>"
+ let ?m' = "nat \<bar>m\<bar>"
+ let ?n' = "nat \<bar>n\<bar>"
+ from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
+ unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
+ from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
+ unfolding zgcd_def by (simp only: zdvd_int)
+ then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
+ then show "k dvd zgcd m n" by simp
+qed
+
+lemma div_zgcd_relprime:
+ assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
+ shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
+proof -
+ from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith
+ let ?g = "zgcd a b"
+ let ?a' = "a div ?g"
+ let ?b' = "b div ?g"
+ let ?g' = "zgcd ?a' ?b'"
+ have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
+ have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
+ from dvdg dvdg' obtain ka kb ka' kb' where
+ kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
+ unfolding dvd_def by blast
+ then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
+ then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+ by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
+ zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
+ have "?g \<noteq> 0" using nz by simp
+ then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
+ from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+ with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
+ with zgcd_pos show "?g' = 1" by simp
+qed
+
+lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
+ by (simp add: zgcd_def abs_if)
+
+lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
+ by (simp add: zgcd_def abs_if)
+
+lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
+ apply (frule_tac b = n and a = m in pos_mod_sign)
+ apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
+ apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
+ apply (frule_tac a = m in pos_mod_bound)
+ apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
+ done
+
+lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
+ apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
+ apply (auto simp add: linorder_neq_iff zgcd_non_0)
+ apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
+ done
+
+lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
+ by (simp add: zgcd_def abs_if)
+
+lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
+ by (simp add: zgcd_def abs_if)
+
+lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
+ by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
+
+lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
+ by (simp add: zgcd_def gcd_1_left)
+
+lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
+ by (simp add: zgcd_def gcd_assoc)
+
+lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
+ apply (rule zgcd_commute [THEN trans])
+ apply (rule zgcd_assoc [THEN trans])
+ apply (rule zgcd_commute [THEN arg_cong])
+ done
+
+lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
+ -- {* addition is an AC-operator *}
+
+lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
+ by (simp del: minus_mult_right [symmetric]
+ add: minus_mult_right nat_mult_distrib zgcd_def abs_if
+ mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
+
+lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
+ by (simp add: abs_if zgcd_zmult_distrib2)
+
+lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
+ by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
+
+lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
+ by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
+
+lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
+ by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
+
+
+definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
+
+lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
+by(simp add:zlcm_def dvd_int_iff)
+
+lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
+by(simp add:zlcm_def dvd_int_iff)
+
+
+lemma dvd_imp_dvd_zlcm1:
+ assumes "k dvd i" shows "k dvd (zlcm i j)"
+proof -
+ have "nat(abs k) dvd nat(abs i)" using `k dvd i`
+ by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
+ thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+qed
+
+lemma dvd_imp_dvd_zlcm2:
+ assumes "k dvd j" shows "k dvd (zlcm i j)"
+proof -
+ have "nat(abs k) dvd nat(abs j)" using `k dvd j`
+ by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
+ thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+qed
+
+
+lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
+by (case_tac "d <0", simp_all)
+
+lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
+by (case_tac "d<0", simp_all)
+
+(* lcm a b is positive for positive a and b *)
+
+lemma lcm_pos:
+ assumes mpos: "m > 0"
+ and npos: "n>0"
+ shows "lcm m n > 0"
+proof(rule ccontr, simp add: lcm_def gcd_zero)
+assume h:"m*n div gcd m n = 0"
+from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
+hence gcdp: "gcd m n > 0" by simp
+with h
+have "m*n < gcd m n"
+ by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
+moreover
+have "gcd m n dvd m" by simp
+ with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
+ with npos have t1:"gcd m n *n \<le> m*n" by simp
+ have "gcd m n \<le> gcd m n*n" using npos by simp
+ with t1 have "gcd m n \<le> m*n" by arith
+ultimately show "False" by simp
+qed
+
+lemma zlcm_pos:
+ assumes anz: "a \<noteq> 0"
+ and bnz: "b \<noteq> 0"
+ shows "0 < zlcm a b"
+proof-
+ let ?na = "nat (abs a)"
+ let ?nb = "nat (abs b)"
+ have nap: "?na >0" using anz by simp
+ have nbp: "?nb >0" using bnz by simp
+ have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
+ thus ?thesis by (simp add: zlcm_def)
+qed
+
+lemma zgcd_code [code]:
+ "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
+ by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
+
+end
--- a/src/HOL/Library/Library.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Library/Library.thy Tue Jun 23 15:48:55 2009 +0100
@@ -25,6 +25,7 @@
Fin_Fun
Float
Formal_Power_Series
+ Fraction_Field
FrechetDeriv
FuncSet
Fundamental_Theorem_Algebra
--- a/src/HOL/Library/Primes.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Library/Primes.thy Tue Jun 23 15:48:55 2009 +0100
@@ -6,9 +6,11 @@
header {* Primality on nat *}
theory Primes
-imports Complex_Main
+imports Complex_Main Legacy_GCD
begin
+hide (open) const GCD.gcd GCD.coprime GCD.prime
+
definition
coprime :: "nat => nat => bool" where
"coprime m n \<longleftrightarrow> gcd m n = 1"
--- a/src/HOL/List.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/List.thy Tue Jun 23 15:48:55 2009 +0100
@@ -363,7 +363,7 @@
val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
$ NilC;
val cs = Syntax.const "_case2" $ case1 $ case2
- val ft = DatatypeCase.case_tr false DatatypePackage.datatype_of_constr
+ val ft = DatatypeCase.case_tr false Datatype.datatype_of_constr
ctxt [x, cs]
in lambda x ft end;
--- a/src/HOL/MetisExamples/Abstraction.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/MetisExamples/Abstraction.thy Tue Jun 23 15:48:55 2009 +0100
@@ -201,7 +201,7 @@
"(cl,f) \<in> CLF ==>
CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==>
f \<in> pset cl \<rightarrow> pset cl"
-by auto
+by fast
(*??no longer terminates, with combinators
by (metis Collect_mem_eq SigmaD2 subsetD)
*)
--- a/src/HOL/Nat.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nat.thy Tue Jun 23 15:48:55 2009 +0100
@@ -849,17 +849,20 @@
lemma less_Suc_induct:
assumes less: "i < j"
and step: "!!i. P i (Suc i)"
- and trans: "!!i j k. P i j ==> P j k ==> P i k"
+ and trans: "!!i j k. i < j ==> j < k ==> P i j ==> P j k ==> P i k"
shows "P i j"
proof -
- from less obtain k where j: "j = Suc(i+k)" by (auto dest: less_imp_Suc_add)
+ from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add)
have "P i (Suc (i + k))"
proof (induct k)
case 0
show ?case by (simp add: step)
next
case (Suc k)
- thus ?case by (auto intro: assms)
+ have "0 + i < Suc k + i" by (rule add_less_mono1) simp
+ hence "i < Suc (i + k)" by (simp add: add_commute)
+ from trans[OF this lessI Suc step]
+ show ?case by simp
qed
thus "P i j" by (simp add: j)
qed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NatTransfer.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,533 @@
+(* Title: HOL/Library/NatTransfer.thy
+ Authors: Jeremy Avigad and Amine Chaieb
+
+ Sets up transfer from nats to ints and
+ back.
+*)
+
+
+header {* NatTransfer *}
+
+theory NatTransfer
+imports Main Parity
+uses ("Tools/transfer_data.ML")
+begin
+
+subsection {* A transfer Method between isomorphic domains*}
+
+definition TransferMorphism:: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> bool"
+ where "TransferMorphism a B = True"
+
+use "Tools/transfer_data.ML"
+
+setup TransferData.setup
+
+
+subsection {* Set up transfer from nat to int *}
+
+(* set up transfer direction *)
+
+lemma TransferMorphism_nat_int: "TransferMorphism nat (op <= (0::int))"
+ by (simp add: TransferMorphism_def)
+
+declare TransferMorphism_nat_int[transfer
+ add mode: manual
+ return: nat_0_le
+ labels: natint
+]
+
+(* basic functions and relations *)
+
+lemma transfer_nat_int_numerals:
+ "(0::nat) = nat 0"
+ "(1::nat) = nat 1"
+ "(2::nat) = nat 2"
+ "(3::nat) = nat 3"
+ by auto
+
+definition
+ tsub :: "int \<Rightarrow> int \<Rightarrow> int"
+where
+ "tsub x y = (if x >= y then x - y else 0)"
+
+lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y"
+ by (simp add: tsub_def)
+
+
+lemma transfer_nat_int_functions:
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)"
+ "(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
+ by (auto simp add: eq_nat_nat_iff nat_mult_distrib
+ nat_power_eq nat_div_distrib nat_mod_distrib tsub_def)
+
+lemma transfer_nat_int_function_closures:
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0"
+ "(x::int) >= 0 \<Longrightarrow> x^n >= 0"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
+ "(0::int) >= 0"
+ "(1::int) >= 0"
+ "(2::int) >= 0"
+ "(3::int) >= 0"
+ "int z >= 0"
+ apply (auto simp add: zero_le_mult_iff tsub_def)
+ apply (case_tac "y = 0")
+ apply auto
+ apply (subst pos_imp_zdiv_nonneg_iff, auto)
+ apply (case_tac "y = 0")
+ apply force
+ apply (rule pos_mod_sign)
+ apply arith
+done
+
+lemma transfer_nat_int_relations:
+ "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
+ (nat (x::int) = nat y) = (x = y)"
+ "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
+ (nat (x::int) < nat y) = (x < y)"
+ "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
+ (nat (x::int) <= nat y) = (x <= y)"
+ "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
+ (nat (x::int) dvd nat y) = (x dvd y)"
+ by (auto simp add: zdvd_int even_nat_def)
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_numerals
+ transfer_nat_int_functions
+ transfer_nat_int_function_closures
+ transfer_nat_int_relations
+]
+
+
+(* first-order quantifiers *)
+
+lemma transfer_nat_int_quantifiers:
+ "(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))"
+ "(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))"
+ by (rule all_nat, rule ex_nat)
+
+(* should we restrict these? *)
+lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
+ (ALL x. Q x \<longrightarrow> P x) = (ALL x. Q x \<longrightarrow> P' x)"
+ by auto
+
+lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
+ (EX x. Q x \<and> P x) = (EX x. Q x \<and> P' x)"
+ by auto
+
+declare TransferMorphism_nat_int[transfer add
+ return: transfer_nat_int_quantifiers
+ cong: all_cong ex_cong]
+
+
+(* if *)
+
+lemma nat_if_cong: "(if P then (nat x) else (nat y)) =
+ nat (if P then x else y)"
+ by auto
+
+declare TransferMorphism_nat_int [transfer add return: nat_if_cong]
+
+
+(* operations with sets *)
+
+definition
+ nat_set :: "int set \<Rightarrow> bool"
+where
+ "nat_set S = (ALL x:S. x >= 0)"
+
+lemma transfer_nat_int_set_functions:
+ "card A = card (int ` A)"
+ "{} = nat ` ({}::int set)"
+ "A Un B = nat ` (int ` A Un int ` B)"
+ "A Int B = nat ` (int ` A Int int ` B)"
+ "{x. P x} = nat ` {x. x >= 0 & P(nat x)}"
+ "{..n} = nat ` {0..int n}"
+ "{m..n} = nat ` {int m..int n}" (* need all variants of these! *)
+ apply (rule card_image [symmetric])
+ apply (auto simp add: inj_on_def image_def)
+ apply (rule_tac x = "int x" in bexI)
+ apply auto
+ apply (rule_tac x = "int x" in bexI)
+ apply auto
+ apply (rule_tac x = "int x" in bexI)
+ apply auto
+ apply (rule_tac x = "int x" in exI)
+ apply auto
+ apply (rule_tac x = "int x" in bexI)
+ apply auto
+ apply (rule_tac x = "int x" in bexI)
+ apply auto
+done
+
+lemma transfer_nat_int_set_function_closures:
+ "nat_set {}"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
+ "x >= 0 \<Longrightarrow> nat_set {x..y}"
+ "nat_set {x. x >= 0 & P x}"
+ "nat_set (int ` C)"
+ "nat_set A \<Longrightarrow> x : A \<Longrightarrow> x >= 0" (* does it hurt to turn this on? *)
+ unfolding nat_set_def apply auto
+done
+
+lemma transfer_nat_int_set_relations:
+ "(finite A) = (finite (int ` A))"
+ "(x : A) = (int x : int ` A)"
+ "(A = B) = (int ` A = int ` B)"
+ "(A < B) = (int ` A < int ` B)"
+ "(A <= B) = (int ` A <= int ` B)"
+
+ apply (rule iffI)
+ apply (erule finite_imageI)
+ apply (erule finite_imageD)
+ apply (auto simp add: image_def expand_set_eq inj_on_def)
+ apply (drule_tac x = "int x" in spec, auto)
+ apply (drule_tac x = "int x" in spec, auto)
+ apply (drule_tac x = "int x" in spec, auto)
+done
+
+lemma transfer_nat_int_set_return_embed: "nat_set A \<Longrightarrow>
+ (int ` nat ` A = A)"
+ by (auto simp add: nat_set_def image_def)
+
+lemma transfer_nat_int_set_cong: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow>
+ {(x::int). x >= 0 & P x} = {x. x >= 0 & P' x}"
+ by auto
+
+declare TransferMorphism_nat_int[transfer add
+ return: transfer_nat_int_set_functions
+ transfer_nat_int_set_function_closures
+ transfer_nat_int_set_relations
+ transfer_nat_int_set_return_embed
+ cong: transfer_nat_int_set_cong
+]
+
+
+(* setsum and setprod *)
+
+(* this handles the case where the *domain* of f is nat *)
+lemma transfer_nat_int_sum_prod:
+ "setsum f A = setsum (%x. f (nat x)) (int ` A)"
+ "setprod f A = setprod (%x. f (nat x)) (int ` A)"
+ apply (subst setsum_reindex)
+ apply (unfold inj_on_def, auto)
+ apply (subst setprod_reindex)
+ apply (unfold inj_on_def o_def, auto)
+done
+
+(* this handles the case where the *range* of f is nat *)
+lemma transfer_nat_int_sum_prod2:
+ "setsum f A = nat(setsum (%x. int (f x)) A)"
+ "setprod f A = nat(setprod (%x. int (f x)) A)"
+ apply (subst int_setsum [symmetric])
+ apply auto
+ apply (subst int_setprod [symmetric])
+ apply auto
+done
+
+lemma transfer_nat_int_sum_prod_closure:
+ "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
+ "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
+ unfolding nat_set_def
+ apply (rule setsum_nonneg)
+ apply auto
+ apply (rule setprod_nonneg)
+ apply auto
+done
+
+(* this version doesn't work, even with nat_set A \<Longrightarrow>
+ x : A \<Longrightarrow> x >= 0 turned on. Why not?
+
+ also: what does =simp=> do?
+
+lemma transfer_nat_int_sum_prod_closure:
+ "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
+ "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
+ unfolding nat_set_def simp_implies_def
+ apply (rule setsum_nonneg)
+ apply auto
+ apply (rule setprod_nonneg)
+ apply auto
+done
+*)
+
+(* Making A = B in this lemma doesn't work. Why not?
+ Also, why aren't setsum_cong and setprod_cong enough,
+ with the previously mentioned rule turned on? *)
+
+lemma transfer_nat_int_sum_prod_cong:
+ "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
+ setsum f A = setsum g B"
+ "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
+ setprod f A = setprod g B"
+ unfolding nat_set_def
+ apply (subst setsum_cong, assumption)
+ apply auto [2]
+ apply (subst setprod_cong, assumption, auto)
+done
+
+declare TransferMorphism_nat_int[transfer add
+ return: transfer_nat_int_sum_prod transfer_nat_int_sum_prod2
+ transfer_nat_int_sum_prod_closure
+ cong: transfer_nat_int_sum_prod_cong]
+
+(* lists *)
+
+definition
+ embed_list :: "nat list \<Rightarrow> int list"
+where
+ "embed_list l = map int l";
+
+definition
+ nat_list :: "int list \<Rightarrow> bool"
+where
+ "nat_list l = nat_set (set l)";
+
+definition
+ return_list :: "int list \<Rightarrow> nat list"
+where
+ "return_list l = map nat l";
+
+thm nat_0_le;
+
+lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
+ embed_list (return_list l) = l";
+ unfolding embed_list_def return_list_def nat_list_def nat_set_def
+ apply (induct l);
+ apply auto;
+done;
+
+lemma transfer_nat_int_list_functions:
+ "l @ m = return_list (embed_list l @ embed_list m)"
+ "[] = return_list []";
+ unfolding return_list_def embed_list_def;
+ apply auto;
+ apply (induct l, auto);
+ apply (induct m, auto);
+done;
+
+(*
+lemma transfer_nat_int_fold1: "fold f l x =
+ fold (%x. f (nat x)) (embed_list l) x";
+*)
+
+
+
+
+subsection {* Set up transfer from int to nat *}
+
+(* set up transfer direction *)
+
+lemma TransferMorphism_int_nat: "TransferMorphism int (UNIV :: nat set)"
+ by (simp add: TransferMorphism_def)
+
+declare TransferMorphism_int_nat[transfer add
+ mode: manual
+(* labels: int-nat *)
+ return: nat_int
+]
+
+
+(* basic functions and relations *)
+
+definition
+ is_nat :: "int \<Rightarrow> bool"
+where
+ "is_nat x = (x >= 0)"
+
+lemma transfer_int_nat_numerals:
+ "0 = int 0"
+ "1 = int 1"
+ "2 = int 2"
+ "3 = int 3"
+ by auto
+
+lemma transfer_int_nat_functions:
+ "(int x) + (int y) = int (x + y)"
+ "(int x) * (int y) = int (x * y)"
+ "tsub (int x) (int y) = int (x - y)"
+ "(int x)^n = int (x^n)"
+ "(int x) div (int y) = int (x div y)"
+ "(int x) mod (int y) = int (x mod y)"
+ by (auto simp add: int_mult tsub_def int_power zdiv_int zmod_int)
+
+lemma transfer_int_nat_function_closures:
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)"
+ "is_nat x \<Longrightarrow> is_nat (x^n)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
+ "is_nat 0"
+ "is_nat 1"
+ "is_nat 2"
+ "is_nat 3"
+ "is_nat (int z)"
+ by (simp_all only: is_nat_def transfer_nat_int_function_closures)
+
+lemma transfer_int_nat_relations:
+ "(int x = int y) = (x = y)"
+ "(int x < int y) = (x < y)"
+ "(int x <= int y) = (x <= y)"
+ "(int x dvd int y) = (x dvd y)"
+ "(even (int x)) = (even x)"
+ by (auto simp add: zdvd_int even_nat_def)
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_numerals
+ transfer_int_nat_functions
+ transfer_int_nat_function_closures
+ transfer_int_nat_relations
+ UNIV_code
+]
+
+
+(* first-order quantifiers *)
+
+lemma transfer_int_nat_quantifiers:
+ "(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))"
+ "(EX (x::int) >= 0. P x) = (EX (x::nat). P (int x))"
+ apply (subst all_nat)
+ apply auto [1]
+ apply (subst ex_nat)
+ apply auto
+done
+
+declare TransferMorphism_int_nat[transfer add
+ return: transfer_int_nat_quantifiers]
+
+
+(* if *)
+
+lemma int_if_cong: "(if P then (int x) else (int y)) =
+ int (if P then x else y)"
+ by auto
+
+declare TransferMorphism_int_nat [transfer add return: int_if_cong]
+
+
+
+(* operations with sets *)
+
+lemma transfer_int_nat_set_functions:
+ "nat_set A \<Longrightarrow> card A = card (nat ` A)"
+ "{} = int ` ({}::nat set)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Int B = int ` (nat ` A Int nat ` B)"
+ "{x. x >= 0 & P x} = int ` {x. P(int x)}"
+ "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
+ (* need all variants of these! *)
+ by (simp_all only: is_nat_def transfer_nat_int_set_functions
+ transfer_nat_int_set_function_closures
+ transfer_nat_int_set_return_embed nat_0_le
+ cong: transfer_nat_int_set_cong)
+
+lemma transfer_int_nat_set_function_closures:
+ "nat_set {}"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
+ "is_nat x \<Longrightarrow> nat_set {x..y}"
+ "nat_set {x. x >= 0 & P x}"
+ "nat_set (int ` C)"
+ "nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x"
+ by (simp_all only: transfer_nat_int_set_function_closures is_nat_def)
+
+lemma transfer_int_nat_set_relations:
+ "nat_set A \<Longrightarrow> finite A = finite (nat ` A)"
+ "is_nat x \<Longrightarrow> nat_set A \<Longrightarrow> (x : A) = (nat x : nat ` A)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A = B) = (nat ` A = nat ` B)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A < B) = (nat ` A < nat ` B)"
+ "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A <= B) = (nat ` A <= nat ` B)"
+ by (simp_all only: is_nat_def transfer_nat_int_set_relations
+ transfer_nat_int_set_return_embed nat_0_le)
+
+lemma transfer_int_nat_set_return_embed: "nat ` int ` A = A"
+ by (simp only: transfer_nat_int_set_relations
+ transfer_nat_int_set_function_closures
+ transfer_nat_int_set_return_embed nat_0_le)
+
+lemma transfer_int_nat_set_cong: "(!!x. P x = P' x) \<Longrightarrow>
+ {(x::nat). P x} = {x. P' x}"
+ by auto
+
+declare TransferMorphism_int_nat[transfer add
+ return: transfer_int_nat_set_functions
+ transfer_int_nat_set_function_closures
+ transfer_int_nat_set_relations
+ transfer_int_nat_set_return_embed
+ cong: transfer_int_nat_set_cong
+]
+
+
+(* setsum and setprod *)
+
+(* this handles the case where the *domain* of f is int *)
+lemma transfer_int_nat_sum_prod:
+ "nat_set A \<Longrightarrow> setsum f A = setsum (%x. f (int x)) (nat ` A)"
+ "nat_set A \<Longrightarrow> setprod f A = setprod (%x. f (int x)) (nat ` A)"
+ apply (subst setsum_reindex)
+ apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff)
+ apply (subst setprod_reindex)
+ apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff
+ cong: setprod_cong)
+done
+
+(* this handles the case where the *range* of f is int *)
+lemma transfer_int_nat_sum_prod2:
+ "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> setsum f A = int(setsum (%x. nat (f x)) A)"
+ "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow>
+ setprod f A = int(setprod (%x. nat (f x)) A)"
+ unfolding is_nat_def
+ apply (subst int_setsum, auto)
+ apply (subst int_setprod, auto simp add: cong: setprod_cong)
+done
+
+declare TransferMorphism_int_nat[transfer add
+ return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2
+ cong: setsum_cong setprod_cong]
+
+
+subsection {* Test it out *}
+
+(* nat to int *)
+
+lemma ex1: "(x::nat) + y = y + x"
+ by auto
+
+thm ex1 [transferred]
+
+lemma ex2: "(a::nat) div b * b + a mod b = a"
+ by (rule mod_div_equality)
+
+thm ex2 [transferred]
+
+lemma ex3: "ALL (x::nat). ALL y. EX z. z >= x + y"
+ by auto
+
+thm ex3 [transferred natint]
+
+lemma ex4: "(x::nat) >= y \<Longrightarrow> (x - y) + y = x"
+ by auto
+
+thm ex4 [transferred]
+
+lemma ex5: "(2::nat) * (SUM i <= n. i) = n * (n + 1)"
+ by (induct n rule: nat_induct, auto)
+
+thm ex5 [transferred]
+
+theorem ex6: "0 <= (n::int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
+ by (rule ex5 [transferred])
+
+thm ex6 [transferred]
+
+thm ex5 [transferred, transferred]
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/Binomial.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,540 @@
+(* Title: Binomial.thy
+ Authors: Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow
+
+
+Defines factorial and the "choose" function, and establishes basic properties.
+
+The original theory "Binomial" was by Lawrence C. Paulson, based on
+the work of Andy Gordon and Florian Kammueller. The approach here,
+which derives the definition of binomial coefficients in terms of the
+factorial function, is due to Jeremy Avigad. The binomial theorem was
+formalized by Tobias Nipkow.
+
+*)
+
+
+header {* Binomial *}
+
+theory Binomial
+imports Cong
+begin
+
+
+subsection {* Main definitions *}
+
+class binomial =
+
+fixes
+ fact :: "'a \<Rightarrow> 'a" and
+ binomial :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "choose" 65)
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: binomial
+
+begin
+
+fun
+ fact_nat :: "nat \<Rightarrow> nat"
+where
+ "fact_nat x =
+ (if x = 0 then 1 else
+ x * fact(x - 1))"
+
+fun
+ binomial_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+ "binomial_nat n k =
+ (if k = 0 then 1 else
+ if n = 0 then 0 else
+ (binomial (n - 1) k) + (binomial (n - 1) (k - 1)))"
+
+instance proof qed
+
+end
+
+(* definitions for the integers *)
+
+instantiation int :: binomial
+
+begin
+
+definition
+ fact_int :: "int \<Rightarrow> int"
+where
+ "fact_int x = (if x >= 0 then int (fact (nat x)) else 0)"
+
+definition
+ binomial_int :: "int => int \<Rightarrow> int"
+where
+ "binomial_int n k = (if n \<ge> 0 \<and> k \<ge> 0 then int (binomial (nat n) (nat k))
+ else 0)"
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_binomial:
+ "(x::int) >= 0 \<Longrightarrow> fact (nat x) = nat (fact x)"
+ "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow> binomial (nat n) (nat k) =
+ nat (binomial n k)"
+ unfolding fact_int_def binomial_int_def
+ by auto
+
+
+lemma transfer_nat_int_binomial_closures:
+ "x >= (0::int) \<Longrightarrow> fact x >= 0"
+ "n >= (0::int) \<Longrightarrow> k >= 0 \<Longrightarrow> binomial n k >= 0"
+ by (auto simp add: fact_int_def binomial_int_def)
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_binomial transfer_nat_int_binomial_closures]
+
+lemma transfer_int_nat_binomial:
+ "fact (int x) = int (fact x)"
+ "binomial (int n) (int k) = int (binomial n k)"
+ unfolding fact_int_def binomial_int_def by auto
+
+lemma transfer_int_nat_binomial_closures:
+ "is_nat x \<Longrightarrow> fact x >= 0"
+ "is_nat n \<Longrightarrow> is_nat k \<Longrightarrow> binomial n k >= 0"
+ by (auto simp add: fact_int_def binomial_int_def)
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_binomial transfer_int_nat_binomial_closures]
+
+
+subsection {* Factorial *}
+
+lemma nat_fact_zero [simp]: "fact (0::nat) = 1"
+ by simp
+
+lemma int_fact_zero [simp]: "fact (0::int) = 1"
+ by (simp add: fact_int_def)
+
+lemma nat_fact_one [simp]: "fact (1::nat) = 1"
+ by simp
+
+lemma nat_fact_Suc_0 [simp]: "fact (Suc 0) = Suc 0"
+ by (simp add: One_nat_def)
+
+lemma int_fact_one [simp]: "fact (1::int) = 1"
+ by (simp add: fact_int_def)
+
+lemma nat_fact_plus_one: "fact ((n::nat) + 1) = (n + 1) * fact n"
+ by simp
+
+lemma nat_fact_Suc: "fact (Suc n) = (Suc n) * fact n"
+ by (simp add: One_nat_def)
+
+lemma int_fact_plus_one:
+ assumes "n >= 0"
+ shows "fact ((n::int) + 1) = (n + 1) * fact n"
+
+ using prems by (rule nat_fact_plus_one [transferred])
+
+lemma nat_fact_reduce: "(n::nat) > 0 \<Longrightarrow> fact n = n * fact (n - 1)"
+ by simp
+
+lemma int_fact_reduce:
+ assumes "(n::int) > 0"
+ shows "fact n = n * fact (n - 1)"
+
+ using prems apply (subst tsub_eq [symmetric], auto)
+ apply (rule nat_fact_reduce [transferred])
+ using prems apply auto
+done
+
+declare fact_nat.simps [simp del]
+
+lemma nat_fact_nonzero [simp]: "fact (n::nat) \<noteq> 0"
+ apply (induct n rule: nat_induct')
+ apply (auto simp add: nat_fact_plus_one)
+done
+
+lemma int_fact_nonzero [simp]: "n >= 0 \<Longrightarrow> fact (n::int) ~= 0"
+ by (simp add: fact_int_def)
+
+lemma nat_fact_gt_zero [simp]: "fact (n :: nat) > 0"
+ by (insert nat_fact_nonzero [of n], arith)
+
+lemma int_fact_gt_zero [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) > 0"
+ by (auto simp add: fact_int_def)
+
+lemma nat_fact_ge_one [simp]: "fact (n :: nat) >= 1"
+ by (insert nat_fact_nonzero [of n], arith)
+
+lemma nat_fact_ge_Suc_0 [simp]: "fact (n :: nat) >= Suc 0"
+ by (insert nat_fact_nonzero [of n], arith)
+
+lemma int_fact_ge_one [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) >= 1"
+ apply (auto simp add: fact_int_def)
+ apply (subgoal_tac "1 = int 1")
+ apply (erule ssubst)
+ apply (subst zle_int)
+ apply auto
+done
+
+lemma nat_dvd_fact [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::nat)"
+ apply (induct n rule: nat_induct')
+ apply (auto simp add: nat_fact_plus_one)
+ apply (subgoal_tac "m = n + 1")
+ apply auto
+done
+
+lemma int_dvd_fact [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::int)"
+ apply (case_tac "1 <= n")
+ apply (induct n rule: int_ge_induct)
+ apply (auto simp add: int_fact_plus_one)
+ apply (subgoal_tac "m = i + 1")
+ apply auto
+done
+
+lemma nat_interval_plus_one: "(i::nat) <= j + 1 \<Longrightarrow>
+ {i..j+1} = {i..j} Un {j+1}"
+ by auto
+
+lemma int_interval_plus_one: "(i::int) <= j + 1 \<Longrightarrow> {i..j+1} = {i..j} Un {j+1}"
+ by auto
+
+lemma nat_fact_altdef: "fact (n::nat) = (PROD i:{1..n}. i)"
+ apply (induct n rule: nat_induct')
+ apply force
+ apply (subst nat_fact_plus_one)
+ apply (subst nat_interval_plus_one)
+ apply auto
+done
+
+lemma int_fact_altdef: "n >= 0 \<Longrightarrow> fact (n::int) = (PROD i:{1..n}. i)"
+ apply (induct n rule: int_ge_induct)
+ apply force
+ apply (subst int_fact_plus_one, assumption)
+ apply (subst int_interval_plus_one)
+ apply auto
+done
+
+subsection {* Infinitely many primes *}
+
+lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1"
+proof-
+ have f1: "fact n + 1 \<noteq> 1" using nat_fact_ge_one [of n] by arith
+ from nat_prime_factor [OF f1]
+ obtain p where "prime p" and "p dvd fact n + 1" by auto
+ hence "p \<le> fact n + 1"
+ by (intro dvd_imp_le, auto)
+ {assume "p \<le> n"
+ from `prime p` have "p \<ge> 1"
+ by (cases p, simp_all)
+ with `p <= n` have "p dvd fact n"
+ by (intro nat_dvd_fact)
+ with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n"
+ by (rule nat_dvd_diff)
+ hence "p dvd 1" by simp
+ hence "p <= 1" by auto
+ moreover from `prime p` have "p > 1" by auto
+ ultimately have False by auto}
+ hence "n < p" by arith
+ with `prime p` and `p <= fact n + 1` show ?thesis by auto
+qed
+
+lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
+using next_prime_bound by auto
+
+lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
+proof
+ assume "finite {(p::nat). prime p}"
+ with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
+ by auto
+ then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
+ by auto
+ with bigger_prime [of b] show False by auto
+qed
+
+
+subsection {* Binomial coefficients *}
+
+lemma nat_choose_zero [simp]: "(n::nat) choose 0 = 1"
+ by simp
+
+lemma int_choose_zero [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
+ by (simp add: binomial_int_def)
+
+lemma nat_zero_choose [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
+ by (induct n rule: nat_induct', auto)
+
+lemma int_zero_choose [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
+ unfolding binomial_int_def apply (case_tac "n < 0")
+ apply force
+ apply (simp del: binomial_nat.simps)
+done
+
+lemma nat_choose_reduce: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+ (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
+ by simp
+
+lemma int_choose_reduce: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+ (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
+ unfolding binomial_int_def apply (subst nat_choose_reduce)
+ apply (auto simp del: binomial_nat.simps
+ simp add: nat_diff_distrib)
+done
+
+lemma nat_choose_plus_one: "((n::nat) + 1) choose (k + 1) =
+ (n choose (k + 1)) + (n choose k)"
+ by (simp add: nat_choose_reduce)
+
+lemma nat_choose_Suc: "(Suc n) choose (Suc k) =
+ (n choose (Suc k)) + (n choose k)"
+ by (simp add: nat_choose_reduce One_nat_def)
+
+lemma int_choose_plus_one: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) =
+ (n choose (k + 1)) + (n choose k)"
+ by (simp add: binomial_int_def nat_choose_plus_one nat_add_distrib
+ del: binomial_nat.simps)
+
+declare binomial_nat.simps [simp del]
+
+lemma nat_choose_self [simp]: "((n::nat) choose n) = 1"
+ by (induct n rule: nat_induct', auto simp add: nat_choose_plus_one)
+
+lemma int_choose_self [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
+ by (auto simp add: binomial_int_def)
+
+lemma nat_choose_one [simp]: "(n::nat) choose 1 = n"
+ by (induct n rule: nat_induct', auto simp add: nat_choose_reduce)
+
+lemma int_choose_one [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
+ by (auto simp add: binomial_int_def)
+
+lemma nat_plus_one_choose_self [simp]: "(n::nat) + 1 choose n = n + 1"
+ apply (induct n rule: nat_induct', force)
+ apply (case_tac "n = 0")
+ apply auto
+ apply (subst nat_choose_reduce)
+ apply (auto simp add: One_nat_def)
+ (* natdiff_cancel_numerals introduces Suc *)
+done
+
+lemma nat_Suc_choose_self [simp]: "(Suc n) choose n = Suc n"
+ using nat_plus_one_choose_self by (simp add: One_nat_def)
+
+lemma int_plus_one_choose_self [rule_format, simp]:
+ "(n::int) \<ge> 0 \<longrightarrow> n + 1 choose n = n + 1"
+ by (auto simp add: binomial_int_def nat_add_distrib)
+
+(* bounded quantification doesn't work with the unicode characters? *)
+lemma nat_choose_pos [rule_format]: "ALL k <= (n::nat).
+ ((n::nat) choose k) > 0"
+ apply (induct n rule: nat_induct')
+ apply force
+ apply clarify
+ apply (case_tac "k = 0")
+ apply force
+ apply (subst nat_choose_reduce)
+ apply auto
+done
+
+lemma int_choose_pos: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
+ ((n::int) choose k) > 0"
+ by (auto simp add: binomial_int_def nat_choose_pos)
+
+lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow>
+ (ALL n. P (n + 1) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (k + 1) \<longrightarrow>
+ P (n + 1) (k + 1))) \<longrightarrow> (ALL k <= n. P n k)"
+ apply (induct n rule: nat_induct')
+ apply auto
+ apply (case_tac "k = 0")
+ apply auto
+ apply (case_tac "k = n + 1")
+ apply auto
+ apply (drule_tac x = n in spec) back back
+ apply (drule_tac x = "k - 1" in spec) back back back
+ apply auto
+done
+
+lemma nat_choose_altdef_aux: "(k::nat) \<le> n \<Longrightarrow>
+ fact k * fact (n - k) * (n choose k) = fact n"
+ apply (rule binomial_induct [of _ k n])
+ apply auto
+proof -
+ fix k :: nat and n
+ assume less: "k < n"
+ assume ih1: "fact k * fact (n - k) * (n choose k) = fact n"
+ hence one: "fact (k + 1) * fact (n - k) * (n choose k) = (k + 1) * fact n"
+ by (subst nat_fact_plus_one, auto)
+ assume ih2: "fact (k + 1) * fact (n - (k + 1)) * (n choose (k + 1)) =
+ fact n"
+ with less have "fact (k + 1) * fact ((n - (k + 1)) + 1) *
+ (n choose (k + 1)) = (n - k) * fact n"
+ by (subst (2) nat_fact_plus_one, auto)
+ with less have two: "fact (k + 1) * fact (n - k) * (n choose (k + 1)) =
+ (n - k) * fact n" by simp
+ have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
+ fact (k + 1) * fact (n - k) * (n choose (k + 1)) +
+ fact (k + 1) * fact (n - k) * (n choose k)"
+ by (subst nat_choose_reduce, auto simp add: ring_simps)
+ also note one
+ also note two
+ also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)"
+ apply (subst nat_fact_plus_one)
+ apply (subst left_distrib [symmetric])
+ apply simp
+ done
+ finally show "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
+ fact (n + 1)" .
+qed
+
+lemma nat_choose_altdef: "(k::nat) \<le> n \<Longrightarrow>
+ n choose k = fact n div (fact k * fact (n - k))"
+ apply (frule nat_choose_altdef_aux)
+ apply (erule subst)
+ apply (simp add: mult_ac)
+done
+
+
+lemma int_choose_altdef:
+ assumes "(0::int) <= k" and "k <= n"
+ shows "n choose k = fact n div (fact k * fact (n - k))"
+
+ apply (subst tsub_eq [symmetric], rule prems)
+ apply (rule nat_choose_altdef [transferred])
+ using prems apply auto
+done
+
+lemma nat_choose_dvd: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
+ unfolding dvd_def apply (frule nat_choose_altdef_aux)
+ (* why don't blast and auto get this??? *)
+ apply (rule exI)
+ apply (erule sym)
+done
+
+lemma int_choose_dvd:
+ assumes "(0::int) <= k" and "k <= n"
+ shows "fact k * fact (n - k) dvd fact n"
+
+ apply (subst tsub_eq [symmetric], rule prems)
+ apply (rule nat_choose_dvd [transferred])
+ using prems apply auto
+done
+
+(* generalizes Tobias Nipkow's proof to any commutative semiring *)
+theorem binomial: "(a+b::'a::{comm_ring_1,power})^n =
+ (SUM k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
+proof (induct n rule: nat_induct')
+ show "?P 0" by simp
+next
+ fix n
+ assume ih: "?P n"
+ have decomp: "{0..n+1} = {0} Un {n+1} Un {1..n}"
+ by auto
+ have decomp2: "{0..n} = {0} Un {1..n}"
+ by auto
+ have decomp3: "{1..n+1} = {n+1} Un {1..n}"
+ by auto
+ have "(a+b)^(n+1) =
+ (a+b) * (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+ using ih by (simp add: power_plus_one)
+ also have "... = a*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k)) +
+ b*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+ by (rule distrib)
+ also have "... = (SUM k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
+ (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k+1))"
+ by (subst (1 2) power_plus_one, simp add: setsum_right_distrib mult_ac)
+ also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
+ (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
+ by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le
+ power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
+ also have "... = a^(n+1) + b^(n+1) +
+ (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
+ (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
+ by (simp add: decomp2 decomp3)
+ also have
+ "... = a^(n+1) + b^(n+1) +
+ (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
+ by (auto simp add: ring_simps setsum_addf [symmetric]
+ nat_choose_reduce)
+ also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
+ using decomp by (simp add: ring_simps)
+ finally show "?P (n + 1)" by simp
+qed
+
+lemma set_explicit: "{S. S = T \<and> P S} = (if P T then {T} else {})"
+ by auto
+
+lemma nat_card_subsets [rule_format]:
+ fixes S :: "'a set"
+ assumes "finite S"
+ shows "ALL k. card {T. T \<le> S \<and> card T = k} = card S choose k"
+ (is "?P S")
+using `finite S`
+proof (induct set: finite)
+ show "?P {}" by (auto simp add: set_explicit)
+ next fix x :: "'a" and F
+ assume iassms: "finite F" "x ~: F"
+ assume ih: "?P F"
+ show "?P (insert x F)" (is "ALL k. ?Q k")
+ proof
+ fix k
+ show "card {T. T \<subseteq> (insert x F) \<and> card T = k} =
+ card (insert x F) choose k" (is "?Q k")
+ proof (induct k rule: nat_induct')
+ from iassms have "{T. T \<le> (insert x F) \<and> card T = 0} = {{}}"
+ apply auto
+ apply (subst (asm) card_0_eq)
+ apply (auto elim: finite_subset)
+ done
+ thus "?Q 0"
+ by auto
+ next fix k
+ show "?Q (k + 1)"
+ proof -
+ from iassms have fin: "finite (insert x F)" by auto
+ hence "{ T. T \<subseteq> insert x F \<and> card T = k + 1} =
+ {T. T \<le> F & card T = k + 1} Un
+ {T. T \<le> insert x F & x : T & card T = k + 1}"
+ by (auto intro!: subsetI)
+ with iassms fin have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
+ card ({T. T \<subseteq> F \<and> card T = k + 1}) +
+ card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1})"
+ apply (subst card_Un_disjoint [symmetric])
+ apply auto
+ (* note: nice! Didn't have to say anything here *)
+ done
+ also from ih have "card ({T. T \<subseteq> F \<and> card T = k + 1}) =
+ card F choose (k+1)" by auto
+ also have "card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}) =
+ card ({T. T <= F & card T = k})"
+ proof -
+ let ?f = "%T. T Un {x}"
+ from iassms have "inj_on ?f {T. T <= F & card T = k}"
+ unfolding inj_on_def by (auto intro!: subsetI)
+ hence "card ({T. T <= F & card T = k}) =
+ card(?f ` {T. T <= F & card T = k})"
+ by (rule card_image [symmetric])
+ also from iassms fin have "?f ` {T. T <= F & card T = k} =
+ {T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}"
+ unfolding image_def
+ (* I can't figure out why this next line takes so long *)
+ apply auto
+ apply (frule (1) finite_subset, force)
+ apply (rule_tac x = "xa - {x}" in exI)
+ apply (subst card_Diff_singleton)
+ apply (auto elim: finite_subset)
+ done
+ finally show ?thesis by (rule sym)
+ qed
+ also from ih have "card ({T. T <= F & card T = k}) = card F choose k"
+ by auto
+ finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
+ card F choose (k + 1) + (card F choose k)".
+ with iassms nat_choose_plus_one show ?thesis
+ by auto
+ qed
+ qed
+ qed
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/Cong.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,1093 @@
+(* Title: HOL/Library/Cong.thy
+ ID:
+ Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+ Thomas M. Rasmussen, Jeremy Avigad
+
+
+Defines congruence (notation: [x = y] (mod z)) for natural numbers and
+integers.
+
+This file combines and revises a number of prior developments.
+
+The original theories "GCD" and "Primes" were by Christophe Tabacznyj
+and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
+gcd, lcm, and prime for the natural numbers.
+
+The original theory "IntPrimes" was by Thomas M. Rasmussen, and
+extended gcd, lcm, primes to the integers. Amine Chaieb provided
+another extension of the notions to the integers, and added a number
+of results to "Primes" and "GCD".
+
+The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
+developed the congruence relations on the integers. The notion was
+extended to the natural numbers by Chiaeb. Jeremy Avigad combined
+these, revised and tidied them, made the development uniform for the
+natural numbers and the integers, and added a number of new theorems.
+
+*)
+
+
+header {* Congruence *}
+
+theory Cong
+imports GCD
+begin
+
+subsection {* Turn off One_nat_def *}
+
+lemma nat_induct' [case_names zero plus1, induct type: nat]:
+ "\<lbrakk> P (0::nat); !!n. P n \<Longrightarrow> P (n + 1)\<rbrakk> \<Longrightarrow> P n"
+by (erule nat_induct) (simp add:One_nat_def)
+
+lemma nat_cases [case_names zero plus1, cases type: nat]:
+ "P (0::nat) \<Longrightarrow> (!!n. P (n + 1)) \<Longrightarrow> P n"
+by(metis nat_induct')
+
+lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"
+by (simp add: One_nat_def)
+
+lemma nat_power_eq_one_eq [simp]:
+ "((x::nat)^m = 1) = (m = 0 | x = 1)"
+by (induct m, auto)
+
+lemma card_insert_if' [simp]: "finite A \<Longrightarrow>
+ card (insert x A) = (if x \<in> A then (card A) else (card A) + 1)"
+by (auto simp add: insert_absorb)
+
+(* why wasn't card_insert_if a simp rule? *)
+declare card_insert_disjoint [simp del]
+
+lemma nat_1' [simp]: "nat 1 = 1"
+by simp
+
+(* For those annoying moments where Suc reappears *)
+lemma Suc_remove: "Suc n = n + 1"
+by simp
+
+declare nat_1 [simp del]
+declare add_2_eq_Suc [simp del]
+declare add_2_eq_Suc' [simp del]
+
+
+declare mod_pos_pos_trivial [simp]
+
+
+subsection {* Main definitions *}
+
+class cong =
+
+fixes
+ cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(mod _'))")
+
+begin
+
+abbreviation
+ notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(mod _'))")
+where
+ "notcong x y m == (~cong x y m)"
+
+end
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: cong
+
+begin
+
+definition
+ cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+where
+ "cong_nat x y m = ((x mod m) = (y mod m))"
+
+instance proof qed
+
+end
+
+
+(* definitions for the integers *)
+
+instantiation int :: cong
+
+begin
+
+definition
+ cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"
+where
+ "cong_int x y m = ((x mod m) = (y mod m))"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_cong:
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> m >= 0 \<Longrightarrow>
+ ([(nat x) = (nat y)] (mod (nat m))) = ([x = y] (mod m))"
+ unfolding cong_int_def cong_nat_def
+ apply (auto simp add: nat_mod_distrib [symmetric])
+ apply (subst (asm) eq_nat_nat_iff)
+ apply (case_tac "m = 0", force, rule pos_mod_sign, force)+
+ apply assumption
+done
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_cong]
+
+lemma transfer_int_nat_cong:
+ "[(int x) = (int y)] (mod (int m)) = [x = y] (mod m)"
+ apply (auto simp add: cong_int_def cong_nat_def)
+ apply (auto simp add: zmod_int [symmetric])
+done
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_cong]
+
+
+subsection {* Congruence *}
+
+(* was zcong_0, etc. *)
+lemma nat_cong_0 [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_0 [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_1 [simp, presburger]: "[(a::nat) = b] (mod 1)"
+ by (unfold cong_nat_def, auto)
+
+lemma nat_cong_Suc_0 [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
+ by (unfold cong_nat_def, auto simp add: One_nat_def)
+
+lemma int_cong_1 [simp, presburger]: "[(a::int) = b] (mod 1)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_refl [simp]: "[(k::nat) = k] (mod m)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_refl [simp]: "[(k::int) = k] (mod m)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_sym: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_sym: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_sym_eq: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_sym_eq: "[(a::int) = b] (mod m) = [b = a] (mod m)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_trans [trans]:
+ "[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_trans [trans]:
+ "[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
+ by (unfold cong_int_def, auto)
+
+lemma nat_cong_add:
+ "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
+ apply (unfold cong_nat_def)
+ apply (subst (1 2) mod_add_eq)
+ apply simp
+done
+
+lemma int_cong_add:
+ "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
+ apply (unfold cong_int_def)
+ apply (subst (1 2) mod_add_left_eq)
+ apply (subst (1 2) mod_add_right_eq)
+ apply simp
+done
+
+lemma int_cong_diff:
+ "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"
+ apply (unfold cong_int_def)
+ apply (subst (1 2) mod_diff_eq)
+ apply simp
+done
+
+lemma int_cong_diff_aux:
+ "(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow>
+ [c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"
+ apply (subst (1 2) tsub_eq)
+ apply (auto intro: int_cong_diff)
+done;
+
+lemma nat_cong_diff:
+ assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and
+ "[c = d] (mod m)"
+ shows "[a - c = b - d] (mod m)"
+
+ using prems by (rule int_cong_diff_aux [transferred]);
+
+lemma nat_cong_mult:
+ "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
+ apply (unfold cong_nat_def)
+ apply (subst (1 2) mod_mult_eq)
+ apply simp
+done
+
+lemma int_cong_mult:
+ "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
+ apply (unfold cong_int_def)
+ apply (subst (1 2) zmod_zmult1_eq)
+ apply (subst (1 2) mult_commute)
+ apply (subst (1 2) zmod_zmult1_eq)
+ apply simp
+done
+
+lemma nat_cong_exp: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+ apply (induct k)
+ apply (auto simp add: nat_cong_refl nat_cong_mult)
+done
+
+lemma int_cong_exp: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+ apply (induct k)
+ apply (auto simp add: int_cong_refl int_cong_mult)
+done
+
+lemma nat_cong_setsum [rule_format]:
+ "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
+ [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto intro: nat_cong_add)
+done
+
+lemma int_cong_setsum [rule_format]:
+ "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
+ [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto intro: int_cong_add)
+done
+
+lemma nat_cong_setprod [rule_format]:
+ "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
+ [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto intro: nat_cong_mult)
+done
+
+lemma int_cong_setprod [rule_format]:
+ "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
+ [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
+ apply (case_tac "finite A")
+ apply (induct set: finite)
+ apply (auto intro: int_cong_mult)
+done
+
+lemma nat_cong_scalar: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+ by (rule nat_cong_mult, simp_all)
+
+lemma int_cong_scalar: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+ by (rule int_cong_mult, simp_all)
+
+lemma nat_cong_scalar2: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+ by (rule nat_cong_mult, simp_all)
+
+lemma int_cong_scalar2: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+ by (rule int_cong_mult, simp_all)
+
+lemma nat_cong_mult_self: "[(a::nat) * m = 0] (mod m)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_cong_mult_self: "[(a::int) * m = 0] (mod m)"
+ by (unfold cong_int_def, auto)
+
+lemma int_cong_eq_diff_cong_0: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
+ apply (rule iffI)
+ apply (erule int_cong_diff [of a b m b b, simplified])
+ apply (erule int_cong_add [of "a - b" 0 m b b, simplified])
+done
+
+lemma int_cong_eq_diff_cong_0_aux: "a >= b \<Longrightarrow>
+ [(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"
+ by (subst tsub_eq, assumption, rule int_cong_eq_diff_cong_0)
+
+lemma nat_cong_eq_diff_cong_0:
+ assumes "(a::nat) >= b"
+ shows "[a = b] (mod m) = [a - b = 0] (mod m)"
+
+ using prems by (rule int_cong_eq_diff_cong_0_aux [transferred])
+
+lemma nat_cong_diff_cong_0':
+ "[(x::nat) = y] (mod n) \<longleftrightarrow>
+ (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
+ apply (case_tac "y <= x")
+ apply (frule nat_cong_eq_diff_cong_0 [where m = n])
+ apply auto [1]
+ apply (subgoal_tac "x <= y")
+ apply (frule nat_cong_eq_diff_cong_0 [where m = n])
+ apply (subst nat_cong_sym_eq)
+ apply auto
+done
+
+lemma nat_cong_altdef: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
+ apply (subst nat_cong_eq_diff_cong_0, assumption)
+ apply (unfold cong_nat_def)
+ apply (simp add: dvd_eq_mod_eq_0 [symmetric])
+done
+
+lemma int_cong_altdef: "[(a::int) = b] (mod m) = (m dvd (a - b))"
+ apply (subst int_cong_eq_diff_cong_0)
+ apply (unfold cong_int_def)
+ apply (simp add: dvd_eq_mod_eq_0 [symmetric])
+done
+
+lemma int_cong_abs: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
+ by (simp add: int_cong_altdef)
+
+lemma int_cong_square:
+ "\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>
+ \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
+ apply (simp only: int_cong_altdef)
+ apply (subst int_prime_dvd_mult_eq [symmetric], assumption)
+ (* any way around this? *)
+ apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
+ apply (auto simp add: ring_simps)
+done
+
+lemma int_cong_mult_rcancel:
+ "coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
+ apply (subst (1 2) int_cong_altdef)
+ apply (subst left_diff_distrib [symmetric])
+ apply (rule int_coprime_dvd_mult_iff)
+ apply (subst int_gcd_commute, assumption)
+done
+
+lemma nat_cong_mult_rcancel:
+ assumes "coprime k (m::nat)"
+ shows "[a * k = b * k] (mod m) = [a = b] (mod m)"
+
+ apply (rule int_cong_mult_rcancel [transferred])
+ using prems apply auto
+done
+
+lemma nat_cong_mult_lcancel:
+ "coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"
+ by (simp add: mult_commute nat_cong_mult_rcancel)
+
+lemma int_cong_mult_lcancel:
+ "coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"
+ by (simp add: mult_commute int_cong_mult_rcancel)
+
+(* was zcong_zgcd_zmult_zmod *)
+lemma int_coprime_cong_mult:
+ "[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
+ \<Longrightarrow> [a = b] (mod m * n)"
+ apply (simp only: int_cong_altdef)
+ apply (erule (2) int_divides_mult)
+done
+
+lemma nat_coprime_cong_mult:
+ assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"
+ shows "[a = b] (mod m * n)"
+
+ apply (rule int_coprime_cong_mult [transferred])
+ using prems apply auto
+done
+
+lemma nat_cong_less_imp_eq: "0 \<le> (a::nat) \<Longrightarrow>
+ a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
+ by (auto simp add: cong_nat_def mod_pos_pos_trivial)
+
+lemma int_cong_less_imp_eq: "0 \<le> (a::int) \<Longrightarrow>
+ a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
+ by (auto simp add: cong_int_def mod_pos_pos_trivial)
+
+lemma nat_cong_less_unique:
+ "0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+ apply auto
+ apply (rule_tac x = "a mod m" in exI)
+ apply (unfold cong_nat_def, auto)
+done
+
+lemma int_cong_less_unique:
+ "0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+ apply auto
+ apply (rule_tac x = "a mod m" in exI)
+ apply (unfold cong_int_def, auto simp add: mod_pos_pos_trivial)
+done
+
+lemma int_cong_iff_lin: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
+ apply (auto simp add: int_cong_altdef dvd_def ring_simps)
+ apply (rule_tac [!] x = "-k" in exI, auto)
+done
+
+lemma nat_cong_iff_lin: "([(a::nat) = b] (mod m)) =
+ (\<exists>k1 k2. b + k1 * m = a + k2 * m)"
+ apply (rule iffI)
+ apply (case_tac "b <= a")
+ apply (subst (asm) nat_cong_altdef, assumption)
+ apply (unfold dvd_def, auto)
+ apply (rule_tac x = k in exI)
+ apply (rule_tac x = 0 in exI)
+ apply (auto simp add: ring_simps)
+ apply (subst (asm) nat_cong_sym_eq)
+ apply (subst (asm) nat_cong_altdef)
+ apply force
+ apply (unfold dvd_def, auto)
+ apply (rule_tac x = 0 in exI)
+ apply (rule_tac x = k in exI)
+ apply (auto simp add: ring_simps)
+ apply (unfold cong_nat_def)
+ apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
+ apply (erule ssubst)back
+ apply (erule subst)
+ apply auto
+done
+
+lemma int_cong_gcd_eq: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
+ apply (subst (asm) int_cong_iff_lin, auto)
+ apply (subst add_commute)
+ apply (subst (2) int_gcd_commute)
+ apply (subst mult_commute)
+ apply (subst int_gcd_add_mult)
+ apply (rule int_gcd_commute)
+done
+
+lemma nat_cong_gcd_eq:
+ assumes "[(a::nat) = b] (mod m)"
+ shows "gcd a m = gcd b m"
+
+ apply (rule int_cong_gcd_eq [transferred])
+ using prems apply auto
+done
+
+lemma nat_cong_imp_coprime: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
+ coprime b m"
+ by (auto simp add: nat_cong_gcd_eq)
+
+lemma int_cong_imp_coprime: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
+ coprime b m"
+ by (auto simp add: int_cong_gcd_eq)
+
+lemma nat_cong_cong_mod: "[(a::nat) = b] (mod m) =
+ [a mod m = b mod m] (mod m)"
+ by (auto simp add: cong_nat_def)
+
+lemma int_cong_cong_mod: "[(a::int) = b] (mod m) =
+ [a mod m = b mod m] (mod m)"
+ by (auto simp add: cong_int_def)
+
+lemma int_cong_minus [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
+ by (subst (1 2) int_cong_altdef, auto)
+
+lemma nat_cong_zero [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
+ by (auto simp add: cong_nat_def)
+
+lemma int_cong_zero [iff]: "[(a::int) = b] (mod 0) = (a = b)"
+ by (auto simp add: cong_int_def)
+
+(*
+lemma int_mod_dvd_mod:
+ "0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"
+ apply (unfold dvd_def, auto)
+ apply (rule mod_mod_cancel)
+ apply auto
+done
+
+lemma mod_dvd_mod:
+ assumes "0 < (m::nat)" and "m dvd b"
+ shows "(a mod b mod m) = (a mod m)"
+
+ apply (rule int_mod_dvd_mod [transferred])
+ using prems apply auto
+done
+*)
+
+lemma nat_cong_add_lcancel:
+ "[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: nat_cong_iff_lin)
+
+lemma int_cong_add_lcancel:
+ "[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: int_cong_iff_lin)
+
+lemma nat_cong_add_rcancel: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: nat_cong_iff_lin)
+
+lemma int_cong_add_rcancel: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: int_cong_iff_lin)
+
+lemma nat_cong_add_lcancel_0: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: nat_cong_iff_lin)
+
+lemma int_cong_add_lcancel_0: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: int_cong_iff_lin)
+
+lemma nat_cong_add_rcancel_0: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: nat_cong_iff_lin)
+
+lemma int_cong_add_rcancel_0: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: int_cong_iff_lin)
+
+lemma nat_cong_dvd_modulus: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
+ [x = y] (mod n)"
+ apply (auto simp add: nat_cong_iff_lin dvd_def)
+ apply (rule_tac x="k1 * k" in exI)
+ apply (rule_tac x="k2 * k" in exI)
+ apply (simp add: ring_simps)
+done
+
+lemma int_cong_dvd_modulus: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
+ [x = y] (mod n)"
+ by (auto simp add: int_cong_altdef dvd_def)
+
+lemma nat_cong_dvd_eq: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+ by (unfold cong_nat_def, auto simp add: dvd_eq_mod_eq_0)
+
+lemma int_cong_dvd_eq: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+ by (unfold cong_int_def, auto simp add: dvd_eq_mod_eq_0)
+
+lemma nat_cong_mod: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
+ by (simp add: cong_nat_def)
+
+lemma int_cong_mod: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
+ by (simp add: cong_int_def)
+
+lemma nat_mod_mult_cong: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0
+ \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
+ by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq)
+
+lemma int_neg_cong: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
+ apply (simp add: int_cong_altdef)
+ apply (subst dvd_minus_iff [symmetric])
+ apply (simp add: ring_simps)
+done
+
+lemma int_cong_modulus_neg: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
+ by (auto simp add: int_cong_altdef)
+
+lemma int_mod_mult_cong: "(a::int) ~= 0 \<Longrightarrow> b ~= 0
+ \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
+ apply (case_tac "b > 0")
+ apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)
+ apply (subst (1 2) int_cong_modulus_neg)
+ apply (unfold cong_int_def)
+ apply (subgoal_tac "a * b = (-a * -b)")
+ apply (erule ssubst)
+ apply (subst zmod_zmult2_eq)
+ apply (auto simp add: mod_add_left_eq)
+done
+
+lemma nat_cong_to_1: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
+ apply (case_tac "a = 0")
+ apply force
+ apply (subst (asm) nat_cong_altdef)
+ apply auto
+done
+
+lemma nat_0_cong_1: "[(0::nat) = 1] (mod n) = (n = 1)"
+ by (unfold cong_nat_def, auto)
+
+lemma int_0_cong_1: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
+ by (unfold cong_int_def, auto simp add: zmult_eq_1_iff)
+
+lemma nat_cong_to_1': "[(a::nat) = 1] (mod n) \<longleftrightarrow>
+ a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
+ apply (case_tac "n = 1")
+ apply auto [1]
+ apply (drule_tac x = "a - 1" in spec)
+ apply force
+ apply (case_tac "a = 0")
+ apply (auto simp add: nat_0_cong_1) [1]
+ apply (rule iffI)
+ apply (drule nat_cong_to_1)
+ apply (unfold dvd_def)
+ apply auto [1]
+ apply (rule_tac x = k in exI)
+ apply (auto simp add: ring_simps) [1]
+ apply (subst nat_cong_altdef)
+ apply (auto simp add: dvd_def)
+done
+
+lemma nat_cong_le: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
+ apply (subst nat_cong_altdef)
+ apply assumption
+ apply (unfold dvd_def, auto simp add: ring_simps)
+ apply (rule_tac x = k in exI)
+ apply auto
+done
+
+lemma nat_cong_solve: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+ apply (case_tac "n = 0")
+ apply force
+ apply (frule nat_bezout [of a n], auto)
+ apply (rule exI, erule ssubst)
+ apply (rule nat_cong_trans)
+ apply (rule nat_cong_add)
+ apply (subst mult_commute)
+ apply (rule nat_cong_mult_self)
+ prefer 2
+ apply simp
+ apply (rule nat_cong_refl)
+ apply (rule nat_cong_refl)
+done
+
+lemma int_cong_solve: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+ apply (case_tac "n = 0")
+ apply (case_tac "a \<ge> 0")
+ apply auto
+ apply (rule_tac x = "-1" in exI)
+ apply auto
+ apply (insert int_bezout [of a n], auto)
+ apply (rule exI)
+ apply (erule subst)
+ apply (rule int_cong_trans)
+ prefer 2
+ apply (rule int_cong_add)
+ apply (rule int_cong_refl)
+ apply (rule int_cong_sym)
+ apply (rule int_cong_mult_self)
+ apply simp
+ apply (subst mult_commute)
+ apply (rule int_cong_refl)
+done
+
+lemma nat_cong_solve_dvd:
+ assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"
+ shows "EX x. [a * x = d] (mod n)"
+proof -
+ from nat_cong_solve [OF a] obtain x where
+ "[a * x = gcd a n](mod n)"
+ by auto
+ hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
+ by (elim nat_cong_scalar2)
+ also from b have "(d div gcd a n) * gcd a n = d"
+ by (rule dvd_div_mult_self)
+ also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
+ by auto
+ finally show ?thesis
+ by auto
+qed
+
+lemma int_cong_solve_dvd:
+ assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"
+ shows "EX x. [a * x = d] (mod n)"
+proof -
+ from int_cong_solve [OF a] obtain x where
+ "[a * x = gcd a n](mod n)"
+ by auto
+ hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
+ by (elim int_cong_scalar2)
+ also from b have "(d div gcd a n) * gcd a n = d"
+ by (rule dvd_div_mult_self)
+ also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
+ by auto
+ finally show ?thesis
+ by auto
+qed
+
+lemma nat_cong_solve_coprime: "coprime (a::nat) n \<Longrightarrow>
+ EX x. [a * x = 1] (mod n)"
+ apply (case_tac "a = 0")
+ apply force
+ apply (frule nat_cong_solve [of a n])
+ apply auto
+done
+
+lemma int_cong_solve_coprime: "coprime (a::int) n \<Longrightarrow>
+ EX x. [a * x = 1] (mod n)"
+ apply (case_tac "a = 0")
+ apply auto
+ apply (case_tac "n \<ge> 0")
+ apply auto
+ apply (subst cong_int_def, auto)
+ apply (frule int_cong_solve [of a n])
+ apply auto
+done
+
+lemma nat_coprime_iff_invertible: "m > (1::nat) \<Longrightarrow> coprime a m =
+ (EX x. [a * x = 1] (mod m))"
+ apply (auto intro: nat_cong_solve_coprime)
+ apply (unfold cong_nat_def, auto intro: nat_invertible_coprime)
+done
+
+lemma int_coprime_iff_invertible: "m > (1::int) \<Longrightarrow> coprime a m =
+ (EX x. [a * x = 1] (mod m))"
+ apply (auto intro: int_cong_solve_coprime)
+ apply (unfold cong_int_def)
+ apply (auto intro: int_invertible_coprime)
+done
+
+lemma int_coprime_iff_invertible': "m > (1::int) \<Longrightarrow> coprime a m =
+ (EX x. 0 <= x & x < m & [a * x = 1] (mod m))"
+ apply (subst int_coprime_iff_invertible)
+ apply auto
+ apply (auto simp add: cong_int_def)
+ apply (rule_tac x = "x mod m" in exI)
+ apply (auto simp add: mod_mult_right_eq [symmetric])
+done
+
+
+lemma nat_cong_cong_lcm: "[(x::nat) = y] (mod a) \<Longrightarrow>
+ [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
+ apply (case_tac "y \<le> x")
+ apply (auto simp add: nat_cong_altdef nat_lcm_least) [1]
+ apply (rule nat_cong_sym)
+ apply (subst (asm) (1 2) nat_cong_sym_eq)
+ apply (auto simp add: nat_cong_altdef nat_lcm_least)
+done
+
+lemma int_cong_cong_lcm: "[(x::int) = y] (mod a) \<Longrightarrow>
+ [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
+ by (auto simp add: int_cong_altdef int_lcm_least) [1]
+
+lemma nat_cong_cong_coprime: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
+ [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
+ apply (frule (1) nat_cong_cong_lcm)back
+ apply (simp add: lcm_nat_def)
+done
+
+lemma int_cong_cong_coprime: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
+ [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
+ apply (frule (1) int_cong_cong_lcm)back
+ apply (simp add: int_lcm_altdef int_cong_abs abs_mult [symmetric])
+done
+
+lemma nat_cong_cong_setprod_coprime [rule_format]: "finite A \<Longrightarrow>
+ (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+ (ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>
+ [x = y] (mod (PROD i:A. m i))"
+ apply (induct set: finite)
+ apply auto
+ apply (rule nat_cong_cong_coprime)
+ apply (subst nat_gcd_commute)
+ apply (rule nat_setprod_coprime)
+ apply auto
+done
+
+lemma int_cong_cong_setprod_coprime [rule_format]: "finite A \<Longrightarrow>
+ (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+ (ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>
+ [x = y] (mod (PROD i:A. m i))"
+ apply (induct set: finite)
+ apply auto
+ apply (rule int_cong_cong_coprime)
+ apply (subst int_gcd_commute)
+ apply (rule int_setprod_coprime)
+ apply auto
+done
+
+lemma nat_binary_chinese_remainder_aux:
+ assumes a: "coprime (m1::nat) m2"
+ shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
+ [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
+proof -
+ from nat_cong_solve_coprime [OF a]
+ obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
+ by auto
+ from a have b: "coprime m2 m1"
+ by (subst nat_gcd_commute)
+ from nat_cong_solve_coprime [OF b]
+ obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
+ by auto
+ have "[m1 * x1 = 0] (mod m1)"
+ by (subst mult_commute, rule nat_cong_mult_self)
+ moreover have "[m2 * x2 = 0] (mod m2)"
+ by (subst mult_commute, rule nat_cong_mult_self)
+ moreover note one two
+ ultimately show ?thesis by blast
+qed
+
+lemma int_binary_chinese_remainder_aux:
+ assumes a: "coprime (m1::int) m2"
+ shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
+ [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
+proof -
+ from int_cong_solve_coprime [OF a]
+ obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
+ by auto
+ from a have b: "coprime m2 m1"
+ by (subst int_gcd_commute)
+ from int_cong_solve_coprime [OF b]
+ obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
+ by auto
+ have "[m1 * x1 = 0] (mod m1)"
+ by (subst mult_commute, rule int_cong_mult_self)
+ moreover have "[m2 * x2 = 0] (mod m2)"
+ by (subst mult_commute, rule int_cong_mult_self)
+ moreover note one two
+ ultimately show ?thesis by blast
+qed
+
+lemma nat_binary_chinese_remainder:
+ assumes a: "coprime (m1::nat) m2"
+ shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+ from nat_binary_chinese_remainder_aux [OF a] obtain b1 b2
+ where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
+ "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
+ by blast
+ let ?x = "u1 * b1 + u2 * b2"
+ have "[?x = u1 * 1 + u2 * 0] (mod m1)"
+ apply (rule nat_cong_add)
+ apply (rule nat_cong_scalar2)
+ apply (rule `[b1 = 1] (mod m1)`)
+ apply (rule nat_cong_scalar2)
+ apply (rule `[b2 = 0] (mod m1)`)
+ done
+ hence "[?x = u1] (mod m1)" by simp
+ have "[?x = u1 * 0 + u2 * 1] (mod m2)"
+ apply (rule nat_cong_add)
+ apply (rule nat_cong_scalar2)
+ apply (rule `[b1 = 0] (mod m2)`)
+ apply (rule nat_cong_scalar2)
+ apply (rule `[b2 = 1] (mod m2)`)
+ done
+ hence "[?x = u2] (mod m2)" by simp
+ with `[?x = u1] (mod m1)` show ?thesis by blast
+qed
+
+lemma int_binary_chinese_remainder:
+ assumes a: "coprime (m1::int) m2"
+ shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+ from int_binary_chinese_remainder_aux [OF a] obtain b1 b2
+ where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
+ "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
+ by blast
+ let ?x = "u1 * b1 + u2 * b2"
+ have "[?x = u1 * 1 + u2 * 0] (mod m1)"
+ apply (rule int_cong_add)
+ apply (rule int_cong_scalar2)
+ apply (rule `[b1 = 1] (mod m1)`)
+ apply (rule int_cong_scalar2)
+ apply (rule `[b2 = 0] (mod m1)`)
+ done
+ hence "[?x = u1] (mod m1)" by simp
+ have "[?x = u1 * 0 + u2 * 1] (mod m2)"
+ apply (rule int_cong_add)
+ apply (rule int_cong_scalar2)
+ apply (rule `[b1 = 0] (mod m2)`)
+ apply (rule int_cong_scalar2)
+ apply (rule `[b2 = 1] (mod m2)`)
+ done
+ hence "[?x = u2] (mod m2)" by simp
+ with `[?x = u1] (mod m1)` show ?thesis by blast
+qed
+
+lemma nat_cong_modulus_mult: "[(x::nat) = y] (mod m * n) \<Longrightarrow>
+ [x = y] (mod m)"
+ apply (case_tac "y \<le> x")
+ apply (simp add: nat_cong_altdef)
+ apply (erule dvd_mult_left)
+ apply (rule nat_cong_sym)
+ apply (subst (asm) nat_cong_sym_eq)
+ apply (simp add: nat_cong_altdef)
+ apply (erule dvd_mult_left)
+done
+
+lemma int_cong_modulus_mult: "[(x::int) = y] (mod m * n) \<Longrightarrow>
+ [x = y] (mod m)"
+ apply (simp add: int_cong_altdef)
+ apply (erule dvd_mult_left)
+done
+
+lemma nat_cong_less_modulus_unique:
+ "[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
+ by (simp add: cong_nat_def)
+
+lemma nat_binary_chinese_remainder_unique:
+ assumes a: "coprime (m1::nat) m2" and
+ nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
+ shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+ from nat_binary_chinese_remainder [OF a] obtain y where
+ "[y = u1] (mod m1)" and "[y = u2] (mod m2)"
+ by blast
+ let ?x = "y mod (m1 * m2)"
+ from nz have less: "?x < m1 * m2"
+ by auto
+ have one: "[?x = u1] (mod m1)"
+ apply (rule nat_cong_trans)
+ prefer 2
+ apply (rule `[y = u1] (mod m1)`)
+ apply (rule nat_cong_modulus_mult)
+ apply (rule nat_cong_mod)
+ using nz apply auto
+ done
+ have two: "[?x = u2] (mod m2)"
+ apply (rule nat_cong_trans)
+ prefer 2
+ apply (rule `[y = u2] (mod m2)`)
+ apply (subst mult_commute)
+ apply (rule nat_cong_modulus_mult)
+ apply (rule nat_cong_mod)
+ using nz apply auto
+ done
+ have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow>
+ z = ?x"
+ proof (clarify)
+ fix z
+ assume "z < m1 * m2"
+ assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)"
+ have "[?x = z] (mod m1)"
+ apply (rule nat_cong_trans)
+ apply (rule `[?x = u1] (mod m1)`)
+ apply (rule nat_cong_sym)
+ apply (rule `[z = u1] (mod m1)`)
+ done
+ moreover have "[?x = z] (mod m2)"
+ apply (rule nat_cong_trans)
+ apply (rule `[?x = u2] (mod m2)`)
+ apply (rule nat_cong_sym)
+ apply (rule `[z = u2] (mod m2)`)
+ done
+ ultimately have "[?x = z] (mod m1 * m2)"
+ by (auto intro: nat_coprime_cong_mult a)
+ with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"
+ apply (intro nat_cong_less_modulus_unique)
+ apply (auto, erule nat_cong_sym)
+ done
+ qed
+ with less one two show ?thesis
+ by auto
+ qed
+
+lemma nat_chinese_remainder_aux:
+ fixes A :: "'a set" and
+ m :: "'a \<Rightarrow> nat"
+ assumes fin: "finite A" and
+ cop: "ALL i : A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+ shows "EX b. (ALL i : A.
+ [b i = 1] (mod m i) \<and> [b i = 0] (mod (PROD j : A - {i}. m j)))"
+proof (rule finite_set_choice, rule fin, rule ballI)
+ fix i
+ assume "i : A"
+ with cop have "coprime (PROD j : A - {i}. m j) (m i)"
+ by (intro nat_setprod_coprime, auto)
+ hence "EX x. [(PROD j : A - {i}. m j) * x = 1] (mod m i)"
+ by (elim nat_cong_solve_coprime)
+ then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
+ by auto
+ moreover have "[(PROD j : A - {i}. m j) * x = 0]
+ (mod (PROD j : A - {i}. m j))"
+ by (subst mult_commute, rule nat_cong_mult_self)
+ ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0]
+ (mod setprod m (A - {i}))"
+ by blast
+qed
+
+lemma nat_chinese_remainder:
+ fixes A :: "'a set" and
+ m :: "'a \<Rightarrow> nat" and
+ u :: "'a \<Rightarrow> nat"
+ assumes
+ fin: "finite A" and
+ cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+ shows "EX x. (ALL i:A. [x = u i] (mod m i))"
+proof -
+ from nat_chinese_remainder_aux [OF fin cop] obtain b where
+ bprop: "ALL i:A. [b i = 1] (mod m i) \<and>
+ [b i = 0] (mod (PROD j : A - {i}. m j))"
+ by blast
+ let ?x = "SUM i:A. (u i) * (b i)"
+ show "?thesis"
+ proof (rule exI, clarify)
+ fix i
+ assume a: "i : A"
+ show "[?x = u i] (mod m i)"
+ proof -
+ from fin a have "?x = (SUM j:{i}. u j * b j) +
+ (SUM j:A-{i}. u j * b j)"
+ by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)
+ hence "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
+ by auto
+ also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
+ u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
+ apply (rule nat_cong_add)
+ apply (rule nat_cong_scalar2)
+ using bprop a apply blast
+ apply (rule nat_cong_setsum)
+ apply (rule nat_cong_scalar2)
+ using bprop apply auto
+ apply (rule nat_cong_dvd_modulus)
+ apply (drule (1) bspec)
+ apply (erule conjE)
+ apply assumption
+ apply (rule dvd_setprod)
+ using fin a apply auto
+ done
+ finally show ?thesis
+ by simp
+ qed
+ qed
+qed
+
+lemma nat_coprime_cong_prod [rule_format]: "finite A \<Longrightarrow>
+ (ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+ (ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>
+ [x = y] (mod (PROD i:A. m i))"
+ apply (induct set: finite)
+ apply auto
+ apply (erule (1) nat_coprime_cong_mult)
+ apply (subst nat_gcd_commute)
+ apply (rule nat_setprod_coprime)
+ apply auto
+done
+
+lemma nat_chinese_remainder_unique:
+ fixes A :: "'a set" and
+ m :: "'a \<Rightarrow> nat" and
+ u :: "'a \<Rightarrow> nat"
+ assumes
+ fin: "finite A" and
+ nz: "ALL i:A. m i \<noteq> 0" and
+ cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+ shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"
+proof -
+ from nat_chinese_remainder [OF fin cop] obtain y where
+ one: "(ALL i:A. [y = u i] (mod m i))"
+ by blast
+ let ?x = "y mod (PROD i:A. m i)"
+ from fin nz have prodnz: "(PROD i:A. m i) \<noteq> 0"
+ by auto
+ hence less: "?x < (PROD i:A. m i)"
+ by auto
+ have cong: "ALL i:A. [?x = u i] (mod m i)"
+ apply auto
+ apply (rule nat_cong_trans)
+ prefer 2
+ using one apply auto
+ apply (rule nat_cong_dvd_modulus)
+ apply (rule nat_cong_mod)
+ using prodnz apply auto
+ apply (rule dvd_setprod)
+ apply (rule fin)
+ apply assumption
+ done
+ have unique: "ALL z. z < (PROD i:A. m i) \<and>
+ (ALL i:A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"
+ proof (clarify)
+ fix z
+ assume zless: "z < (PROD i:A. m i)"
+ assume zcong: "(ALL i:A. [z = u i] (mod m i))"
+ have "ALL i:A. [?x = z] (mod m i)"
+ apply clarify
+ apply (rule nat_cong_trans)
+ using cong apply (erule bspec)
+ apply (rule nat_cong_sym)
+ using zcong apply auto
+ done
+ with fin cop have "[?x = z] (mod (PROD i:A. m i))"
+ by (intro nat_coprime_cong_prod, auto)
+ with zless less show "z = ?x"
+ apply (intro nat_cong_less_modulus_unique)
+ apply (auto, erule nat_cong_sym)
+ done
+ qed
+ from less cong unique show ?thesis
+ by blast
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/Fib.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,319 @@
+(* Title: Fib.thy
+ Authors: Lawrence C. Paulson, Jeremy Avigad
+
+
+Defines the fibonacci function.
+
+The original "Fib" is due to Lawrence C. Paulson, and was adapted by
+Jeremy Avigad.
+*)
+
+
+header {* Fib *}
+
+theory Fib
+imports Binomial
+begin
+
+
+subsection {* Main definitions *}
+
+class fib =
+
+fixes
+ fib :: "'a \<Rightarrow> 'a"
+
+
+(* definition for the natural numbers *)
+
+instantiation nat :: fib
+
+begin
+
+fun
+ fib_nat :: "nat \<Rightarrow> nat"
+where
+ "fib_nat n =
+ (if n = 0 then 0 else
+ (if n = 1 then 1 else
+ fib (n - 1) + fib (n - 2)))"
+
+instance proof qed
+
+end
+
+(* definition for the integers *)
+
+instantiation int :: fib
+
+begin
+
+definition
+ fib_int :: "int \<Rightarrow> int"
+where
+ "fib_int n = (if n >= 0 then int (fib (nat n)) else 0)"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_fib:
+ "(x::int) >= 0 \<Longrightarrow> fib (nat x) = nat (fib x)"
+ unfolding fib_int_def by auto
+
+lemma transfer_nat_int_fib_closure:
+ "n >= (0::int) \<Longrightarrow> fib n >= 0"
+ by (auto simp add: fib_int_def)
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_fib transfer_nat_int_fib_closure]
+
+lemma transfer_int_nat_fib:
+ "fib (int n) = int (fib n)"
+ unfolding fib_int_def by auto
+
+lemma transfer_int_nat_fib_closure:
+ "is_nat n \<Longrightarrow> fib n >= 0"
+ unfolding fib_int_def by auto
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_fib transfer_int_nat_fib_closure]
+
+
+subsection {* Fibonacci numbers *}
+
+lemma nat_fib_0 [simp]: "fib (0::nat) = 0"
+ by simp
+
+lemma int_fib_0 [simp]: "fib (0::int) = 0"
+ unfolding fib_int_def by simp
+
+lemma nat_fib_1 [simp]: "fib (1::nat) = 1"
+ by simp
+
+lemma nat_fib_Suc_0 [simp]: "fib (Suc 0) = Suc 0"
+ by simp
+
+lemma int_fib_1 [simp]: "fib (1::int) = 1"
+ unfolding fib_int_def by simp
+
+lemma nat_fib_reduce: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+ by simp
+
+declare fib_nat.simps [simp del]
+
+lemma int_fib_reduce: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+ unfolding fib_int_def
+ by (auto simp add: nat_fib_reduce nat_diff_distrib)
+
+lemma int_fib_neg [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
+ unfolding fib_int_def by auto
+
+lemma nat_fib_2 [simp]: "fib (2::nat) = 1"
+ by (subst nat_fib_reduce, auto)
+
+lemma int_fib_2 [simp]: "fib (2::int) = 1"
+ by (subst int_fib_reduce, auto)
+
+lemma nat_fib_plus_2: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
+ by (subst nat_fib_reduce, auto simp add: One_nat_def)
+(* the need for One_nat_def is due to the natdiff_cancel_numerals
+ procedure *)
+
+lemma nat_fib_induct: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow>
+ (!!n. P n \<Longrightarrow> P (n + 1) \<Longrightarrow> P (n + 2)) \<Longrightarrow> P n"
+ apply (atomize, induct n rule: nat_less_induct)
+ apply auto
+ apply (case_tac "n = 0", force)
+ apply (case_tac "n = 1", force)
+ apply (subgoal_tac "n >= 2")
+ apply (frule_tac x = "n - 1" in spec)
+ apply (drule_tac x = "n - 2" in spec)
+ apply (drule_tac x = "n - 2" in spec)
+ apply auto
+ apply (auto simp add: One_nat_def) (* again, natdiff_cancel *)
+done
+
+lemma nat_fib_add: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) +
+ fib k * fib n"
+ apply (induct n rule: nat_fib_induct)
+ apply auto
+ apply (subst nat_fib_reduce)
+ apply (auto simp add: ring_simps)
+ apply (subst (1 3 5) nat_fib_reduce)
+ apply (auto simp add: ring_simps Suc_remove)
+(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
+ apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
+ apply (erule ssubst) back back
+ apply (erule ssubst) back
+ apply auto
+done
+
+lemma nat_fib_add': "fib (n + Suc k) = fib (Suc k) * fib (Suc n) +
+ fib k * fib n"
+ using nat_fib_add by (auto simp add: One_nat_def)
+
+
+(* transfer from nats to ints *)
+lemma int_fib_add [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
+ fib (n + k + 1) = fib (k + 1) * fib (n + 1) +
+ fib k * fib n "
+
+ by (rule nat_fib_add [transferred])
+
+lemma nat_fib_neq_0: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
+ apply (induct n rule: nat_fib_induct)
+ apply (auto simp add: nat_fib_plus_2)
+done
+
+lemma nat_fib_gr_0: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
+ by (frule nat_fib_neq_0, simp)
+
+lemma int_fib_gr_0: "(n::int) > 0 \<Longrightarrow> fib n > 0"
+ unfolding fib_int_def by (simp add: nat_fib_gr_0)
+
+text {*
+ \medskip Concrete Mathematics, page 278: Cassini's identity. The proof is
+ much easier using integers, not natural numbers!
+*}
+
+lemma int_fib_Cassini_aux: "fib (int n + 2) * fib (int n) -
+ (fib (int n + 1))^2 = (-1)^(n + 1)"
+ apply (induct n)
+ apply (auto simp add: ring_simps power2_eq_square int_fib_reduce
+ power_add)
+done
+
+lemma int_fib_Cassini: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n -
+ (fib (n + 1))^2 = (-1)^(nat n + 1)"
+ by (insert int_fib_Cassini_aux [of "nat n"], auto)
+
+(*
+lemma int_fib_Cassini': "n >= 0 \<Longrightarrow> fib (n + 2) * fib n =
+ (fib (n + 1))^2 + (-1)^(nat n + 1)"
+ by (frule int_fib_Cassini, simp)
+*)
+
+lemma int_fib_Cassini': "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
+ (if even n then tsub ((fib (n + 1))^2) 1
+ else (fib (n + 1))^2 + 1)"
+ apply (frule int_fib_Cassini, auto simp add: pos_int_even_equiv_nat_even)
+ apply (subst tsub_eq)
+ apply (insert int_fib_gr_0 [of "n + 1"], force)
+ apply auto
+done
+
+lemma nat_fib_Cassini: "fib ((n::nat) + 2) * fib n =
+ (if even n then (fib (n + 1))^2 - 1
+ else (fib (n + 1))^2 + 1)"
+
+ by (rule int_fib_Cassini' [transferred, of n], auto)
+
+
+text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
+
+lemma nat_coprime_fib_plus_1: "coprime (fib (n::nat)) (fib (n + 1))"
+ apply (induct n rule: nat_fib_induct)
+ apply auto
+ apply (subst (2) nat_fib_reduce)
+ apply (auto simp add: Suc_remove) (* again, natdiff_cancel *)
+ apply (subst add_commute, auto)
+ apply (subst nat_gcd_commute, auto simp add: ring_simps)
+done
+
+lemma nat_coprime_fib_Suc: "coprime (fib n) (fib (Suc n))"
+ using nat_coprime_fib_plus_1 by (simp add: One_nat_def)
+
+lemma int_coprime_fib_plus_1:
+ "n >= 0 \<Longrightarrow> coprime (fib (n::int)) (fib (n + 1))"
+ by (erule nat_coprime_fib_plus_1 [transferred])
+
+lemma nat_gcd_fib_add: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
+ apply (simp add: nat_gcd_commute [of "fib m"])
+ apply (rule nat_cases [of _ m])
+ apply simp
+ apply (subst add_assoc [symmetric])
+ apply (simp add: nat_fib_add)
+ apply (subst nat_gcd_commute)
+ apply (subst mult_commute)
+ apply (subst nat_gcd_add_mult)
+ apply (subst nat_gcd_commute)
+ apply (rule nat_gcd_mult_cancel)
+ apply (rule nat_coprime_fib_plus_1)
+done
+
+lemma int_gcd_fib_add [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+ gcd (fib (m::int)) (fib (n + m)) = gcd (fib m) (fib n)"
+ by (erule nat_gcd_fib_add [transferred])
+
+lemma nat_gcd_fib_diff: "(m::nat) \<le> n \<Longrightarrow>
+ gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+ by (simp add: nat_gcd_fib_add [symmetric, of _ "n-m"])
+
+lemma int_gcd_fib_diff: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow>
+ gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+ by (simp add: int_gcd_fib_add [symmetric, of _ "n-m"])
+
+lemma nat_gcd_fib_mod: "0 < (m::nat) \<Longrightarrow>
+ gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+proof (induct n rule: less_induct)
+ case (less n)
+ from less.prems have pos_m: "0 < m" .
+ show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+ proof (cases "m < n")
+ case True note m_n = True
+ then have m_n': "m \<le> n" by auto
+ with pos_m have pos_n: "0 < n" by auto
+ with pos_m m_n have diff: "n - m < n" by auto
+ have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
+ by (simp add: mod_if [of n]) (insert m_n, auto)
+ also have "\<dots> = gcd (fib m) (fib (n - m))"
+ by (simp add: less.hyps diff pos_m)
+ also have "\<dots> = gcd (fib m) (fib n)" by (simp add: nat_gcd_fib_diff m_n')
+ finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
+ next
+ case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+ by (cases "m = n") auto
+ qed
+qed
+
+lemma int_gcd_fib_mod:
+ assumes "0 < (m::int)" and "0 <= n"
+ shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+
+ apply (rule nat_gcd_fib_mod [transferred])
+ using prems apply auto
+done
+
+lemma nat_fib_gcd: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"
+ -- {* Law 6.111 *}
+ apply (induct m n rule: nat_gcd_induct)
+ apply (simp_all add: nat_gcd_non_0 nat_gcd_commute nat_gcd_fib_mod)
+done
+
+lemma int_fib_gcd: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+ fib (gcd (m::int) n) = gcd (fib m) (fib n)"
+ by (erule nat_fib_gcd [transferred])
+
+lemma nat_atMost_plus_one: "{..(k::nat) + 1} = insert (k + 1) {..k}"
+ by auto
+
+theorem nat_fib_mult_eq_setsum:
+ "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+ apply (induct n)
+ apply (auto simp add: nat_atMost_plus_one nat_fib_plus_2 ring_simps)
+done
+
+theorem nat_fib_mult_eq_setsum':
+ "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+ using nat_fib_mult_eq_setsum by (simp add: One_nat_def)
+
+theorem int_fib_mult_eq_setsum [rule_format]:
+ "n >= 0 \<Longrightarrow> fib ((n::int) + 1) * fib n = (\<Sum>k \<in> {0..n}. fib k * fib k)"
+ by (erule nat_fib_mult_eq_setsum [transferred])
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/MiscAlgebra.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,355 @@
+(* Title: MiscAlgebra.thy
+ Author: Jeremy Avigad
+
+These are things that can be added to the Algebra library.
+*)
+
+theory MiscAlgebra
+imports
+ "~~/src/HOL/Algebra/Ring"
+ "~~/src/HOL/Algebra/FiniteProduct"
+begin;
+
+(* finiteness stuff *)
+
+lemma int_bounded_set1 [intro]: "finite {(x::int). a < x & x < b & P x}"
+ apply (subgoal_tac "{x. a < x & x < b & P x} <= {a<..<b}")
+ apply (erule finite_subset)
+ apply auto
+done
+
+
+(* The rest is for the algebra libraries *)
+
+(* These go in Group.thy. *)
+
+(*
+ Show that the units in any monoid give rise to a group.
+
+ The file Residues.thy provides some infrastructure to use
+ facts about the unit group within the ring locale.
+*)
+
+
+constdefs
+ units_of :: "('a, 'b) monoid_scheme => 'a monoid"
+ "units_of G == (| carrier = Units G,
+ Group.monoid.mult = Group.monoid.mult G,
+ one = one G |)";
+
+(*
+
+lemma (in monoid) Units_mult_closed [intro]:
+ "x : Units G ==> y : Units G ==> x \<otimes> y : Units G"
+ apply (unfold Units_def)
+ apply (clarsimp)
+ apply (rule_tac x = "xaa \<otimes> xa" in bexI)
+ apply auto
+ apply (subst m_assoc)
+ apply auto
+ apply (subst (2) m_assoc [symmetric])
+ apply auto
+ apply (subst m_assoc)
+ apply auto
+ apply (subst (2) m_assoc [symmetric])
+ apply auto
+done
+
+*)
+
+lemma (in monoid) units_group: "group(units_of G)"
+ apply (unfold units_of_def)
+ apply (rule groupI)
+ apply auto
+ apply (subst m_assoc)
+ apply auto
+ apply (rule_tac x = "inv x" in bexI)
+ apply auto
+done
+
+lemma (in comm_monoid) units_comm_group: "comm_group(units_of G)"
+ apply (rule group.group_comm_groupI)
+ apply (rule units_group)
+ apply (insert prems)
+ apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
+ apply auto;
+done;
+
+lemma units_of_carrier: "carrier (units_of G) = Units G"
+ by (unfold units_of_def, auto)
+
+lemma units_of_mult: "mult(units_of G) = mult G"
+ by (unfold units_of_def, auto)
+
+lemma units_of_one: "one(units_of G) = one G"
+ by (unfold units_of_def, auto)
+
+lemma (in monoid) units_of_inv: "x : Units G ==>
+ m_inv (units_of G) x = m_inv G x"
+ apply (rule sym)
+ apply (subst m_inv_def)
+ apply (rule the1_equality)
+ apply (rule ex_ex1I)
+ apply (subst (asm) Units_def)
+ apply auto
+ apply (erule inv_unique)
+ apply auto
+ apply (rule Units_closed)
+ apply (simp_all only: units_of_carrier [symmetric])
+ apply (insert units_group)
+ apply auto
+ apply (subst units_of_mult [symmetric])
+ apply (subst units_of_one [symmetric])
+ apply (erule group.r_inv, assumption)
+ apply (subst units_of_mult [symmetric])
+ apply (subst units_of_one [symmetric])
+ apply (erule group.l_inv, assumption)
+done
+
+lemma (in group) inj_on_const_mult: "a: (carrier G) ==>
+ inj_on (%x. a \<otimes> x) (carrier G)"
+ by (unfold inj_on_def, auto)
+
+lemma (in group) surj_const_mult: "a : (carrier G) ==>
+ (%x. a \<otimes> x) ` (carrier G) = (carrier G)"
+ apply (auto simp add: image_def)
+ apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
+ apply auto
+(* auto should get this. I suppose we need "comm_monoid_simprules"
+ for mult_ac rewriting. *)
+ apply (subst m_assoc [symmetric])
+ apply auto
+done
+
+lemma (in group) l_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+ (x \<otimes> a = x) = (a = one G)"
+ apply auto
+ apply (subst l_cancel [symmetric])
+ prefer 4
+ apply (erule ssubst)
+ apply auto
+done
+
+lemma (in group) r_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+ (a \<otimes> x = x) = (a = one G)"
+ apply auto
+ apply (subst r_cancel [symmetric])
+ prefer 4
+ apply (erule ssubst)
+ apply auto
+done
+
+(* Is there a better way to do this? *)
+
+lemma (in group) l_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+ (x = x \<otimes> a) = (a = one G)"
+ by (subst eq_commute, simp)
+
+lemma (in group) r_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+ (x = a \<otimes> x) = (a = one G)"
+ by (subst eq_commute, simp)
+
+(* This should be generalized to arbitrary groups, not just commutative
+ ones, using Lagrange's theorem. *)
+
+lemma (in comm_group) power_order_eq_one:
+ assumes fin [simp]: "finite (carrier G)"
+ and a [simp]: "a : carrier G"
+ shows "a (^) card(carrier G) = one G"
+proof -
+ have "(\<Otimes>x:carrier G. x) = (\<Otimes>x:carrier G. a \<otimes> x)"
+ by (subst (2) finprod_reindex [symmetric],
+ auto simp add: Pi_def inj_on_const_mult surj_const_mult)
+ also have "\<dots> = (\<Otimes>x:carrier G. a) \<otimes> (\<Otimes>x:carrier G. x)"
+ by (auto simp add: finprod_multf Pi_def)
+ also have "(\<Otimes>x:carrier G. a) = a (^) card(carrier G)"
+ by (auto simp add: finprod_const)
+ finally show ?thesis
+(* uses the preceeding lemma *)
+ by auto
+qed
+
+
+(* Miscellaneous *)
+
+lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> ~= \<one>\<^bsub>R\<^esub> \<Longrightarrow> ALL x : carrier R - {\<zero>\<^bsub>R\<^esub>}.
+ x : Units R \<Longrightarrow> field R"
+ apply (unfold_locales)
+ apply (insert prems, auto)
+ apply (rule trans)
+ apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
+ apply assumption
+ apply (subst m_assoc)
+ apply (auto simp add: Units_r_inv)
+ apply (unfold Units_def)
+ apply auto
+done
+
+lemma (in monoid) inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
+ x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
+ apply (subgoal_tac "x : Units G")
+ apply (subgoal_tac "y = inv x \<otimes> \<one>")
+ apply simp
+ apply (erule subst)
+ apply (subst m_assoc [symmetric])
+ apply auto
+ apply (unfold Units_def)
+ apply auto
+done
+
+lemma (in comm_monoid) comm_inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
+ x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
+ apply (rule inv_char)
+ apply auto
+ apply (subst m_comm, auto)
+done
+
+lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"
+ apply (rule inv_char)
+ apply (auto simp add: l_minus r_minus)
+done
+
+lemma (in monoid) inv_eq_imp_eq: "x : Units G \<Longrightarrow> y : Units G \<Longrightarrow>
+ inv x = inv y \<Longrightarrow> x = y"
+ apply (subgoal_tac "inv(inv x) = inv(inv y)")
+ apply (subst (asm) Units_inv_inv)+
+ apply auto
+done
+
+lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> : Units R"
+ apply (unfold Units_def)
+ apply auto
+ apply (rule_tac x = "\<ominus> \<one>" in bexI)
+ apply auto
+ apply (simp add: l_minus r_minus)
+done
+
+lemma (in monoid) inv_one [simp]: "inv \<one> = \<one>"
+ apply (rule inv_char)
+ apply auto
+done
+
+lemma (in ring) inv_eq_neg_one_eq: "x : Units R \<Longrightarrow> (inv x = \<ominus> \<one>) = (x = \<ominus> \<one>)"
+ apply auto
+ apply (subst Units_inv_inv [symmetric])
+ apply auto
+done
+
+lemma (in monoid) inv_eq_one_eq: "x : Units G \<Longrightarrow> (inv x = \<one>) = (x = \<one>)"
+ apply auto
+ apply (subst Units_inv_inv [symmetric])
+ apply auto
+done
+
+
+(* This goes in FiniteProduct *)
+
+lemma (in comm_monoid) finprod_UN_disjoint:
+ "finite I \<Longrightarrow> (ALL i:I. finite (A i)) \<longrightarrow> (ALL i:I. ALL j:I. i ~= j \<longrightarrow>
+ (A i) Int (A j) = {}) \<longrightarrow>
+ (ALL i:I. ALL x: (A i). g x : carrier G) \<longrightarrow>
+ finprod G g (UNION I A) = finprod G (%i. finprod G g (A i)) I"
+ apply (induct set: finite)
+ apply force
+ apply clarsimp
+ apply (subst finprod_Un_disjoint)
+ apply blast
+ apply (erule finite_UN_I)
+ apply blast
+ apply (fastsimp)
+ apply (auto intro!: funcsetI finprod_closed)
+done
+
+lemma (in comm_monoid) finprod_Union_disjoint:
+ "[| finite C; (ALL A:C. finite A & (ALL x:A. f x : carrier G));
+ (ALL A:C. ALL B:C. A ~= B --> A Int B = {}) |]
+ ==> finprod G f (Union C) = finprod G (finprod G f) C"
+ apply (frule finprod_UN_disjoint [of C id f])
+ apply (unfold Union_def id_def, auto)
+done
+
+lemma (in comm_monoid) finprod_one [rule_format]:
+ "finite A \<Longrightarrow> (ALL x:A. f x = \<one>) \<longrightarrow>
+ finprod G f A = \<one>"
+by (induct set: finite) auto
+
+
+(* need better simplification rules for rings *)
+(* the next one holds more generally for abelian groups *)
+
+lemma (in cring) sum_zero_eq_neg:
+ "x : carrier R \<Longrightarrow> y : carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
+ apply (subgoal_tac "\<ominus> y = \<zero> \<oplus> \<ominus> y")
+ apply (erule ssubst)back
+ apply (erule subst)
+ apply (simp add: ring_simprules)+
+done
+
+(* there's a name conflict -- maybe "domain" should be
+ "integral_domain" *)
+
+lemma (in Ring.domain) square_eq_one:
+ fixes x
+ assumes [simp]: "x : carrier R" and
+ "x \<otimes> x = \<one>"
+ shows "x = \<one> | x = \<ominus>\<one>"
+proof -
+ have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
+ by (simp add: ring_simprules)
+ also with `x \<otimes> x = \<one>` have "\<dots> = \<zero>"
+ by (simp add: ring_simprules)
+ finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
+ hence "(x \<oplus> \<one>) = \<zero> | (x \<oplus> \<ominus> \<one>) = \<zero>"
+ by (intro integral, auto)
+ thus ?thesis
+ apply auto
+ apply (erule notE)
+ apply (rule sum_zero_eq_neg)
+ apply auto
+ apply (subgoal_tac "x = \<ominus> (\<ominus> \<one>)")
+ apply (simp add: ring_simprules)
+ apply (rule sum_zero_eq_neg)
+ apply auto
+ done
+qed
+
+lemma (in Ring.domain) inv_eq_self: "x : Units R \<Longrightarrow>
+ x = inv x \<Longrightarrow> x = \<one> | x = \<ominus> \<one>"
+ apply (rule square_eq_one)
+ apply auto
+ apply (erule ssubst)back
+ apply (erule Units_r_inv)
+done
+
+
+(*
+ The following translates theorems about groups to the facts about
+ the units of a ring. (The list should be expanded as more things are
+ needed.)
+*)
+
+lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow>
+ finite (Units R)"
+ by (rule finite_subset, auto)
+
+(* this belongs with MiscAlgebra.thy *)
+lemma (in monoid) units_of_pow:
+ "x : Units G \<Longrightarrow> x (^)\<^bsub>units_of G\<^esub> (n::nat) = x (^)\<^bsub>G\<^esub> n"
+ apply (induct n)
+ apply (auto simp add: units_group group.is_monoid
+ monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult
+ One_nat_def)
+done
+
+lemma (in cring) units_power_order_eq_one: "finite (Units R) \<Longrightarrow> a : Units R
+ \<Longrightarrow> a (^) card(Units R) = \<one>"
+ apply (subst units_of_carrier [symmetric])
+ apply (subst units_of_one [symmetric])
+ apply (subst units_of_pow [symmetric])
+ apply assumption
+ apply (rule comm_group.power_order_eq_one)
+ apply (rule units_comm_group)
+ apply (unfold units_of_def, auto)
+done
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/ROOT.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,1 @@
+use_thys ["Fib","Residues"];
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/Residues.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,469 @@
+(* Title: HOL/Library/Residues.thy
+ ID:
+ Author: Jeremy Avigad
+
+ An algebraic treatment of residue rings, and resulting proofs of
+ Euler's theorem and Wilson's theorem.
+*)
+
+header {* Residue rings *}
+
+theory Residues
+imports
+ UniqueFactorization
+ Binomial
+ MiscAlgebra
+begin
+
+
+(*
+
+ A locale for residue rings
+
+*)
+
+constdefs
+ residue_ring :: "int => int ring"
+ "residue_ring m == (|
+ carrier = {0..m - 1},
+ mult = (%x y. (x * y) mod m),
+ one = 1,
+ zero = 0,
+ add = (%x y. (x + y) mod m) |)"
+
+locale residues =
+ fixes m :: int and R (structure)
+ assumes m_gt_one: "m > 1"
+ defines "R == residue_ring m"
+
+context residues begin
+
+lemma abelian_group: "abelian_group R"
+ apply (insert m_gt_one)
+ apply (rule abelian_groupI)
+ apply (unfold R_def residue_ring_def)
+ apply (auto simp add: mod_pos_pos_trivial mod_add_right_eq [symmetric]
+ add_ac)
+ apply (case_tac "x = 0")
+ apply force
+ apply (subgoal_tac "(x + (m - x)) mod m = 0")
+ apply (erule bexI)
+ apply auto
+done
+
+lemma comm_monoid: "comm_monoid R"
+ apply (insert m_gt_one)
+ apply (unfold R_def residue_ring_def)
+ apply (rule comm_monoidI)
+ apply auto
+ apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m")
+ apply (erule ssubst)
+ apply (subst zmod_zmult1_eq [symmetric])+
+ apply (simp_all only: mult_ac)
+done
+
+lemma cring: "cring R"
+ apply (rule cringI)
+ apply (rule abelian_group)
+ apply (rule comm_monoid)
+ apply (unfold R_def residue_ring_def, auto)
+ apply (subst mod_add_eq [symmetric])
+ apply (subst mult_commute)
+ apply (subst zmod_zmult1_eq [symmetric])
+ apply (simp add: ring_simps)
+done
+
+end
+
+sublocale residues < cring
+ by (rule cring)
+
+
+context residues begin
+
+(* These lemmas translate back and forth between internal and
+ external concepts *)
+
+lemma res_carrier_eq: "carrier R = {0..m - 1}"
+ by (unfold R_def residue_ring_def, auto)
+
+lemma res_add_eq: "x \<oplus> y = (x + y) mod m"
+ by (unfold R_def residue_ring_def, auto)
+
+lemma res_mult_eq: "x \<otimes> y = (x * y) mod m"
+ by (unfold R_def residue_ring_def, auto)
+
+lemma res_zero_eq: "\<zero> = 0"
+ by (unfold R_def residue_ring_def, auto)
+
+lemma res_one_eq: "\<one> = 1"
+ by (unfold R_def residue_ring_def units_of_def residue_ring_def, auto)
+
+lemma res_units_eq: "Units R = { x. 0 < x & x < m & coprime x m}"
+ apply (insert m_gt_one)
+ apply (unfold Units_def R_def residue_ring_def)
+ apply auto
+ apply (subgoal_tac "x ~= 0")
+ apply auto
+ apply (rule int_invertible_coprime)
+ apply (subgoal_tac "x ~= 0")
+ apply auto
+ apply (subst (asm) int_coprime_iff_invertible')
+ apply (rule m_gt_one)
+ apply (auto simp add: cong_int_def mult_commute)
+done
+
+lemma res_neg_eq: "\<ominus> x = (- x) mod m"
+ apply (insert m_gt_one)
+ apply (unfold R_def a_inv_def m_inv_def residue_ring_def)
+ apply auto
+ apply (rule the_equality)
+ apply auto
+ apply (subst mod_add_right_eq [symmetric])
+ apply auto
+ apply (subst mod_add_left_eq [symmetric])
+ apply auto
+ apply (subgoal_tac "y mod m = - x mod m")
+ apply simp
+ apply (subst zmod_eq_dvd_iff)
+ apply auto
+done
+
+lemma finite [iff]: "finite(carrier R)"
+ by (subst res_carrier_eq, auto)
+
+lemma finite_Units [iff]: "finite(Units R)"
+ by (subst res_units_eq, auto)
+
+(* The function a -> a mod m maps the integers to the
+ residue classes. The following lemmas show that this mapping
+ respects addition and multiplication on the integers. *)
+
+lemma mod_in_carrier [iff]: "a mod m : carrier R"
+ apply (unfold res_carrier_eq)
+ apply (insert m_gt_one, auto)
+done
+
+lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m"
+ by (unfold R_def residue_ring_def, auto, arith)
+
+lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m"
+ apply (unfold R_def residue_ring_def, auto)
+ apply (subst zmod_zmult1_eq [symmetric])
+ apply (subst mult_commute)
+ apply (subst zmod_zmult1_eq [symmetric])
+ apply (subst mult_commute)
+ apply auto
+done
+
+lemma zero_cong: "\<zero> = 0"
+ apply (unfold R_def residue_ring_def, auto)
+done
+
+lemma one_cong: "\<one> = 1 mod m"
+ apply (insert m_gt_one)
+ apply (unfold R_def residue_ring_def, auto)
+done
+
+(* revise algebra library to use 1? *)
+lemma pow_cong: "(x mod m) (^) n = x^n mod m"
+ apply (insert m_gt_one)
+ apply (induct n)
+ apply (auto simp add: nat_pow_def one_cong One_nat_def)
+ apply (subst mult_commute)
+ apply (rule mult_cong)
+done
+
+lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m"
+ apply (rule sym)
+ apply (rule sum_zero_eq_neg)
+ apply auto
+ apply (subst add_cong)
+ apply (subst zero_cong)
+ apply auto
+done
+
+lemma (in residues) prod_cong:
+ "finite A \<Longrightarrow> (\<Otimes> i:A. (f i) mod m) = (PROD i:A. f i) mod m"
+ apply (induct set: finite)
+ apply (auto simp: one_cong mult_cong)
+done
+
+lemma (in residues) sum_cong:
+ "finite A \<Longrightarrow> (\<Oplus> i:A. (f i) mod m) = (SUM i: A. f i) mod m"
+ apply (induct set: finite)
+ apply (auto simp: zero_cong add_cong)
+done
+
+lemma mod_in_res_units [simp]: "1 < m \<Longrightarrow> coprime a m \<Longrightarrow>
+ a mod m : Units R"
+ apply (subst res_units_eq, auto)
+ apply (insert pos_mod_sign [of m a])
+ apply (subgoal_tac "a mod m ~= 0")
+ apply arith
+ apply auto
+ apply (subst (asm) int_gcd_commute)
+ apply (subst (asm) int_gcd_mult)
+ apply auto
+ apply (subst (asm) int_gcd_red)
+ apply (subst int_gcd_commute, assumption)
+done
+
+lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))"
+ unfolding cong_int_def by auto
+
+(* Simplifying with these will translate a ring equation in R to a
+ congruence. *)
+
+lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong
+ prod_cong sum_cong neg_cong res_eq_to_cong
+
+(* Other useful facts about the residue ring *)
+
+lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2"
+ apply (simp add: res_one_eq res_neg_eq)
+ apply (insert m_gt_one)
+ apply (subgoal_tac "~(m > 2)")
+ apply arith
+ apply (rule notI)
+ apply (subgoal_tac "-1 mod m = m - 1")
+ apply force
+ apply (subst mod_add_self2 [symmetric])
+ apply (subst mod_pos_pos_trivial)
+ apply auto
+done
+
+end
+
+
+(* prime residues *)
+
+locale residues_prime =
+ fixes p :: int and R (structure)
+ assumes p_prime [intro]: "prime p"
+ defines "R == residue_ring p"
+
+sublocale residues_prime < residues p
+ apply (unfold R_def residues_def)
+ using p_prime apply auto
+done
+
+context residues_prime begin
+
+lemma is_field: "field R"
+ apply (rule cring.field_intro2)
+ apply (rule cring)
+ apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq
+ res_units_eq)
+ apply (rule classical)
+ apply (erule notE)
+ apply (subst int_gcd_commute)
+ apply (rule int_prime_imp_coprime)
+ apply (rule p_prime)
+ apply (rule notI)
+ apply (frule zdvd_imp_le)
+ apply auto
+done
+
+lemma res_prime_units_eq: "Units R = {1..p - 1}"
+ apply (subst res_units_eq)
+ apply auto
+ apply (subst int_gcd_commute)
+ apply (rule int_prime_imp_coprime)
+ apply (rule p_prime)
+ apply (rule zdvd_not_zless)
+ apply auto
+done
+
+end
+
+sublocale residues_prime < field
+ by (rule is_field)
+
+
+(*
+ Test cases: Euler's theorem and Wilson's theorem.
+*)
+
+
+subsection{* Euler's theorem *}
+
+(* the definition of the phi function *)
+
+constdefs
+ phi :: "int => nat"
+ "phi m == card({ x. 0 < x & x < m & gcd x m = 1})"
+
+lemma phi_zero [simp]: "phi 0 = 0"
+ apply (subst phi_def)
+(* Auto hangs here. Once again, where is the simplification rule
+ 1 == Suc 0 coming from? *)
+ apply (auto simp add: card_eq_0_iff)
+(* Add card_eq_0_iff as a simp rule? delete card_empty_imp? *)
+done
+
+lemma phi_one [simp]: "phi 1 = 0"
+ apply (auto simp add: phi_def card_eq_0_iff)
+done
+
+lemma (in residues) phi_eq: "phi m = card(Units R)"
+ by (simp add: phi_def res_units_eq)
+
+lemma (in residues) euler_theorem1:
+ assumes a: "gcd a m = 1"
+ shows "[a^phi m = 1] (mod m)"
+proof -
+ from a m_gt_one have [simp]: "a mod m : Units R"
+ by (intro mod_in_res_units)
+ from phi_eq have "(a mod m) (^) (phi m) = (a mod m) (^) (card (Units R))"
+ by simp
+ also have "\<dots> = \<one>"
+ by (intro units_power_order_eq_one, auto)
+ finally show ?thesis
+ by (simp add: res_to_cong_simps)
+qed
+
+(* In fact, there is a two line proof!
+
+lemma (in residues) euler_theorem1:
+ assumes a: "gcd a m = 1"
+ shows "[a^phi m = 1] (mod m)"
+proof -
+ have "(a mod m) (^) (phi m) = \<one>"
+ by (simp add: phi_eq units_power_order_eq_one a m_gt_one)
+ thus ?thesis
+ by (simp add: res_to_cong_simps)
+qed
+
+*)
+
+(* outside the locale, we can relax the restriction m > 1 *)
+
+lemma euler_theorem:
+ assumes "m >= 0" and "gcd a m = 1"
+ shows "[a^phi m = 1] (mod m)"
+proof (cases)
+ assume "m = 0 | m = 1"
+ thus ?thesis by auto
+next
+ assume "~(m = 0 | m = 1)"
+ with prems show ?thesis
+ by (intro residues.euler_theorem1, unfold residues_def, auto)
+qed
+
+lemma (in residues_prime) phi_prime: "phi p = (nat p - 1)"
+ apply (subst phi_eq)
+ apply (subst res_prime_units_eq)
+ apply auto
+done
+
+lemma phi_prime: "prime p \<Longrightarrow> phi p = (nat p - 1)"
+ apply (rule residues_prime.phi_prime)
+ apply (erule residues_prime.intro)
+done
+
+lemma fermat_theorem:
+ assumes "prime p" and "~ (p dvd a)"
+ shows "[a^(nat p - 1) = 1] (mod p)"
+proof -
+ from prems have "[a^phi p = 1] (mod p)"
+ apply (intro euler_theorem)
+ (* auto should get this next part. matching across
+ substitutions is needed. *)
+ apply (frule int_prime_gt_1, arith)
+ apply (subst int_gcd_commute, erule int_prime_imp_coprime, assumption)
+ done
+ also have "phi p = nat p - 1"
+ by (rule phi_prime, rule prems)
+ finally show ?thesis .
+qed
+
+
+subsection {* Wilson's theorem *}
+
+lemma (in field) inv_pair_lemma: "x : Units R \<Longrightarrow> y : Units R \<Longrightarrow>
+ {x, inv x} ~= {y, inv y} \<Longrightarrow> {x, inv x} Int {y, inv y} = {}"
+ apply auto
+ apply (erule notE)
+ apply (erule inv_eq_imp_eq)
+ apply auto
+ apply (erule notE)
+ apply (erule inv_eq_imp_eq)
+ apply auto
+done
+
+lemma (in residues_prime) wilson_theorem1:
+ assumes a: "p > 2"
+ shows "[fact (p - 1) = - 1] (mod p)"
+proof -
+ let ?InversePairs = "{ {x, inv x} | x. x : Units R - {\<one>, \<ominus> \<one>}}"
+ have UR: "Units R = {\<one>, \<ominus> \<one>} Un (Union ?InversePairs)"
+ by auto
+ have "(\<Otimes>i: Units R. i) =
+ (\<Otimes>i: {\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i: Union ?InversePairs. i)"
+ apply (subst UR)
+ apply (subst finprod_Un_disjoint)
+ apply (auto intro:funcsetI)
+ apply (drule sym, subst (asm) inv_eq_one_eq)
+ apply auto
+ apply (drule sym, subst (asm) inv_eq_neg_one_eq)
+ apply auto
+ done
+ also have "(\<Otimes>i: {\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
+ apply (subst finprod_insert)
+ apply auto
+ apply (frule one_eq_neg_one)
+ apply (insert a, force)
+ done
+ also have "(\<Otimes>i:(Union ?InversePairs). i) =
+ (\<Otimes> A: ?InversePairs. (\<Otimes> y:A. y))"
+ apply (subst finprod_Union_disjoint)
+ apply force
+ apply force
+ apply clarify
+ apply (rule inv_pair_lemma)
+ apply auto
+ done
+ also have "\<dots> = \<one>"
+ apply (rule finprod_one)
+ apply auto
+ apply (subst finprod_insert)
+ apply auto
+ apply (frule inv_eq_self)
+ apply (auto)
+ done
+ finally have "(\<Otimes>i: Units R. i) = \<ominus> \<one>"
+ by simp
+ also have "(\<Otimes>i: Units R. i) = (\<Otimes>i: Units R. i mod p)"
+ apply (rule finprod_cong')
+ apply (auto)
+ apply (subst (asm) res_prime_units_eq)
+ apply auto
+ done
+ also have "\<dots> = (PROD i: Units R. i) mod p"
+ apply (rule prod_cong)
+ apply auto
+ done
+ also have "\<dots> = fact (p - 1) mod p"
+ apply (subst int_fact_altdef)
+ apply (insert prems, force)
+ apply (subst res_prime_units_eq, rule refl)
+ done
+ finally have "fact (p - 1) mod p = \<ominus> \<one>".
+ thus ?thesis
+ by (simp add: res_to_cong_simps)
+qed
+
+lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)"
+ apply (frule int_prime_gt_1)
+ apply (case_tac "p = 2")
+ apply (subst int_fact_altdef, simp)
+ apply (subst cong_int_def)
+ apply simp
+ apply (rule residues_prime.wilson_theorem1)
+ apply (rule residues_prime.intro)
+ apply auto
+done
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NewNumberTheory/UniqueFactorization.thy Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,967 @@
+(* Title: UniqueFactorization.thy
+ ID:
+ Author: Jeremy Avigad
+
+
+ Unique factorization for the natural numbers and the integers.
+
+ Note: there were previous Isabelle formalizations of unique
+ factorization due to Thomas Marthedal Rasmussen, and, building on
+ that, by Jeremy Avigad and David Gray.
+*)
+
+header {* UniqueFactorization *}
+
+theory UniqueFactorization
+imports Cong Multiset
+begin
+
+(* inherited from Multiset *)
+declare One_nat_def [simp del]
+
+(* As a simp or intro rule,
+
+ prime p \<Longrightarrow> p > 0
+
+ wreaks havoc here. When the premise includes ALL x :# M. prime x, it
+ leads to the backchaining
+
+ x > 0
+ prime x
+ x :# M which is, unfortunately,
+ count M x > 0
+*)
+
+
+(* useful facts *)
+
+lemma setsum_Un2: "finite (A Un B) \<Longrightarrow>
+ setsum f (A Un B) = setsum f (A - B) + setsum f (B - A) +
+ setsum f (A Int B)"
+ apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
+ apply (erule ssubst)
+ apply (subst setsum_Un_disjoint)
+ apply auto
+ apply (subst setsum_Un_disjoint)
+ apply auto
+done
+
+lemma setprod_Un2: "finite (A Un B) \<Longrightarrow>
+ setprod f (A Un B) = setprod f (A - B) * setprod f (B - A) *
+ setprod f (A Int B)"
+ apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
+ apply (erule ssubst)
+ apply (subst setprod_Un_disjoint)
+ apply auto
+ apply (subst setprod_Un_disjoint)
+ apply auto
+done
+
+(* Should this go in Multiset.thy? *)
+(* TN: No longer an intro-rule; needed only once and might get in the way *)
+lemma multiset_eqI: "[| !!x. count M x = count N x |] ==> M = N"
+ by (subst multiset_eq_conv_count_eq, blast)
+
+(* Here is a version of set product for multisets. Is it worth moving
+ to multiset.thy? If so, one should similarly define msetsum for abelian
+ semirings, using of_nat. Also, is it worth developing bounded quantifiers
+ "ALL i :# M. P i"?
+*)
+
+constdefs
+ msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b"
+ "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)"
+
+syntax
+ "_msetprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"
+ ("(3PROD _:#_. _)" [0, 51, 10] 10)
+
+translations
+ "PROD i :# A. b" == "msetprod (%i. b) A"
+
+lemma msetprod_Un: "msetprod f (A+B) = msetprod f A * msetprod f B"
+ apply (simp add: msetprod_def power_add)
+ apply (subst setprod_Un2)
+ apply auto
+ apply (subgoal_tac
+ "(PROD x:set_of A - set_of B. f x ^ count A x * f x ^ count B x) =
+ (PROD x:set_of A - set_of B. f x ^ count A x)")
+ apply (erule ssubst)
+ apply (subgoal_tac
+ "(PROD x:set_of B - set_of A. f x ^ count A x * f x ^ count B x) =
+ (PROD x:set_of B - set_of A. f x ^ count B x)")
+ apply (erule ssubst)
+ apply (subgoal_tac "(PROD x:set_of A. f x ^ count A x) =
+ (PROD x:set_of A - set_of B. f x ^ count A x) *
+ (PROD x:set_of A Int set_of B. f x ^ count A x)")
+ apply (erule ssubst)
+ apply (subgoal_tac "(PROD x:set_of B. f x ^ count B x) =
+ (PROD x:set_of B - set_of A. f x ^ count B x) *
+ (PROD x:set_of A Int set_of B. f x ^ count B x)")
+ apply (erule ssubst)
+ apply (subst setprod_timesf)
+ apply (force simp add: mult_ac)
+ apply (subst setprod_Un_disjoint [symmetric])
+ apply (auto intro: setprod_cong)
+ apply (subst setprod_Un_disjoint [symmetric])
+ apply (auto intro: setprod_cong)
+done
+
+
+subsection {* unique factorization: multiset version *}
+
+lemma multiset_prime_factorization_exists [rule_format]: "n > 0 -->
+ (EX M. (ALL (p::nat) : set_of M. prime p) & n = (PROD i :# M. i))"
+proof (rule nat_less_induct, clarify)
+ fix n :: nat
+ assume ih: "ALL m < n. 0 < m --> (EX M. (ALL p : set_of M. prime p) & m =
+ (PROD i :# M. i))"
+ assume "(n::nat) > 0"
+ then have "n = 1 | (n > 1 & prime n) | (n > 1 & ~ prime n)"
+ by arith
+ moreover
+ {
+ assume "n = 1"
+ then have "(ALL p : set_of {#}. prime p) & n = (PROD i :# {#}. i)"
+ by (auto simp add: msetprod_def)
+ }
+ moreover
+ {
+ assume "n > 1" and "prime n"
+ then have "(ALL p : set_of {# n #}. prime p) & n = (PROD i :# {# n #}. i)"
+ by (auto simp add: msetprod_def)
+ }
+ moreover
+ {
+ assume "n > 1" and "~ prime n"
+ from prems nat_not_prime_eq_prod
+ obtain m k where "n = m * k & 1 < m & m < n & 1 < k & k < n"
+ by blast
+ with ih obtain Q R where "(ALL p : set_of Q. prime p) & m = (PROD i:#Q. i)"
+ and "(ALL p: set_of R. prime p) & k = (PROD i:#R. i)"
+ by blast
+ hence "(ALL p: set_of (Q + R). prime p) & n = (PROD i :# Q + R. i)"
+ by (auto simp add: prems msetprod_Un set_of_union)
+ then have "EX M. (ALL p : set_of M. prime p) & n = (PROD i :# M. i)"..
+ }
+ ultimately show "EX M. (ALL p : set_of M. prime p) & n = (PROD i::nat:#M. i)"
+ by blast
+qed
+
+lemma multiset_prime_factorization_unique_aux:
+ fixes a :: nat
+ assumes "(ALL p : set_of M. prime p)" and
+ "(ALL p : set_of N. prime p)" and
+ "(PROD i :# M. i) dvd (PROD i:# N. i)"
+ shows
+ "count M a <= count N a"
+proof cases
+ assume "a : set_of M"
+ with prems have a: "prime a"
+ by auto
+ with prems have "a ^ count M a dvd (PROD i :# M. i)"
+ by (auto intro: dvd_setprod simp add: msetprod_def)
+ also have "... dvd (PROD i :# N. i)"
+ by (rule prems)
+ also have "... = (PROD i : (set_of N). i ^ (count N i))"
+ by (simp add: msetprod_def)
+ also have "... =
+ a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))"
+ proof (cases)
+ assume "a : set_of N"
+ hence b: "set_of N = {a} Un (set_of N - {a})"
+ by auto
+ thus ?thesis
+ by (subst (1) b, subst setprod_Un_disjoint, auto)
+ next
+ assume "a ~: set_of N"
+ thus ?thesis
+ by auto
+ qed
+ finally have "a ^ count M a dvd
+ a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))".
+ moreover have "coprime (a ^ count M a)
+ (PROD i : (set_of N - {a}). i ^ (count N i))"
+ apply (subst nat_gcd_commute)
+ apply (rule nat_setprod_coprime)
+ apply (rule nat_primes_imp_powers_coprime)
+ apply (insert prems, auto)
+ done
+ ultimately have "a ^ count M a dvd a^(count N a)"
+ by (elim nat_coprime_dvd_mult)
+ with a show ?thesis
+ by (intro power_dvd_imp_le, auto)
+next
+ assume "a ~: set_of M"
+ thus ?thesis by auto
+qed
+
+lemma multiset_prime_factorization_unique:
+ assumes "(ALL (p::nat) : set_of M. prime p)" and
+ "(ALL p : set_of N. prime p)" and
+ "(PROD i :# M. i) = (PROD i:# N. i)"
+ shows
+ "M = N"
+proof -
+ {
+ fix a
+ from prems have "count M a <= count N a"
+ by (intro multiset_prime_factorization_unique_aux, auto)
+ moreover from prems have "count N a <= count M a"
+ by (intro multiset_prime_factorization_unique_aux, auto)
+ ultimately have "count M a = count N a"
+ by auto
+ }
+ thus ?thesis by (simp add:multiset_eq_conv_count_eq)
+qed
+
+constdefs
+ multiset_prime_factorization :: "nat => nat multiset"
+ "multiset_prime_factorization n ==
+ if n > 0 then (THE M. ((ALL p : set_of M. prime p) &
+ n = (PROD i :# M. i)))
+ else {#}"
+
+lemma multiset_prime_factorization: "n > 0 ==>
+ (ALL p : set_of (multiset_prime_factorization n). prime p) &
+ n = (PROD i :# (multiset_prime_factorization n). i)"
+ apply (unfold multiset_prime_factorization_def)
+ apply clarsimp
+ apply (frule multiset_prime_factorization_exists)
+ apply clarify
+ apply (rule theI)
+ apply (insert multiset_prime_factorization_unique, blast)+
+done
+
+
+subsection {* Prime factors and multiplicity for nats and ints *}
+
+class unique_factorization =
+
+fixes
+ multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" and
+ prime_factors :: "'a \<Rightarrow> 'a set"
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: unique_factorization
+
+begin
+
+definition
+ multiplicity_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+ "multiplicity_nat p n = count (multiset_prime_factorization n) p"
+
+definition
+ prime_factors_nat :: "nat \<Rightarrow> nat set"
+where
+ "prime_factors_nat n = set_of (multiset_prime_factorization n)"
+
+instance proof qed
+
+end
+
+(* definitions for the integers *)
+
+instantiation int :: unique_factorization
+
+begin
+
+definition
+ multiplicity_int :: "int \<Rightarrow> int \<Rightarrow> nat"
+where
+ "multiplicity_int p n = multiplicity (nat p) (nat n)"
+
+definition
+ prime_factors_int :: "int \<Rightarrow> int set"
+where
+ "prime_factors_int n = int ` (prime_factors (nat n))"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up transfer *}
+
+lemma transfer_nat_int_prime_factors:
+ "prime_factors (nat n) = nat ` prime_factors n"
+ unfolding prime_factors_int_def apply auto
+ by (subst transfer_int_nat_set_return_embed, assumption)
+
+lemma transfer_nat_int_prime_factors_closure: "n >= 0 \<Longrightarrow>
+ nat_set (prime_factors n)"
+ by (auto simp add: nat_set_def prime_factors_int_def)
+
+lemma transfer_nat_int_multiplicity: "p >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+ multiplicity (nat p) (nat n) = multiplicity p n"
+ by (auto simp add: multiplicity_int_def)
+
+declare TransferMorphism_nat_int[transfer add return:
+ transfer_nat_int_prime_factors transfer_nat_int_prime_factors_closure
+ transfer_nat_int_multiplicity]
+
+
+lemma transfer_int_nat_prime_factors:
+ "prime_factors (int n) = int ` prime_factors n"
+ unfolding prime_factors_int_def by auto
+
+lemma transfer_int_nat_prime_factors_closure: "is_nat n \<Longrightarrow>
+ nat_set (prime_factors n)"
+ by (simp only: transfer_nat_int_prime_factors_closure is_nat_def)
+
+lemma transfer_int_nat_multiplicity:
+ "multiplicity (int p) (int n) = multiplicity p n"
+ by (auto simp add: multiplicity_int_def)
+
+declare TransferMorphism_int_nat[transfer add return:
+ transfer_int_nat_prime_factors transfer_int_nat_prime_factors_closure
+ transfer_int_nat_multiplicity]
+
+
+subsection {* Properties of prime factors and multiplicity for nats and ints *}
+
+lemma int_prime_factors_ge_0 [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
+ by (unfold prime_factors_int_def, auto)
+
+lemma nat_prime_factors_prime [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
+ apply (case_tac "n = 0")
+ apply (simp add: prime_factors_nat_def multiset_prime_factorization_def)
+ apply (auto simp add: prime_factors_nat_def multiset_prime_factorization)
+done
+
+lemma int_prime_factors_prime [intro]:
+ assumes "n >= 0" and "p : prime_factors (n::int)"
+ shows "prime p"
+
+ apply (rule nat_prime_factors_prime [transferred, of n p])
+ using prems apply auto
+done
+
+lemma nat_prime_factors_gt_0 [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
+ by (frule nat_prime_factors_prime, auto)
+
+lemma int_prime_factors_gt_0 [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow>
+ p > (0::int)"
+ by (frule (1) int_prime_factors_prime, auto)
+
+lemma nat_prime_factors_finite [iff]: "finite (prime_factors (n::nat))"
+ by (unfold prime_factors_nat_def, auto)
+
+lemma int_prime_factors_finite [iff]: "finite (prime_factors (n::int))"
+ by (unfold prime_factors_int_def, auto)
+
+lemma nat_prime_factors_altdef: "prime_factors (n::nat) =
+ {p. multiplicity p n > 0}"
+ by (force simp add: prime_factors_nat_def multiplicity_nat_def)
+
+lemma int_prime_factors_altdef: "prime_factors (n::int) =
+ {p. p >= 0 & multiplicity p n > 0}"
+ apply (unfold prime_factors_int_def multiplicity_int_def)
+ apply (subst nat_prime_factors_altdef)
+ apply (auto simp add: image_def)
+done
+
+lemma nat_prime_factorization: "(n::nat) > 0 \<Longrightarrow>
+ n = (PROD p : prime_factors n. p^(multiplicity p n))"
+ by (frule multiset_prime_factorization,
+ simp add: prime_factors_nat_def multiplicity_nat_def msetprod_def)
+
+thm nat_prime_factorization [transferred]
+
+lemma int_prime_factorization:
+ assumes "(n::int) > 0"
+ shows "n = (PROD p : prime_factors n. p^(multiplicity p n))"
+
+ apply (rule nat_prime_factorization [transferred, of n])
+ using prems apply auto
+done
+
+lemma nat_neq_zero_eq_gt_zero: "((x::nat) ~= 0) = (x > 0)"
+ by auto
+
+lemma nat_prime_factorization_unique:
+ "S = { (p::nat) . f p > 0} \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
+ n = (PROD p : S. p^(f p)) \<Longrightarrow>
+ S = prime_factors n & (ALL p. f p = multiplicity p n)"
+ apply (subgoal_tac "multiset_prime_factorization n = Abs_multiset
+ f")
+ apply (unfold prime_factors_nat_def multiplicity_nat_def)
+ apply (simp add: set_of_def count_def Abs_multiset_inverse multiset_def)
+ apply (unfold multiset_prime_factorization_def)
+ apply (subgoal_tac "n > 0")
+ prefer 2
+ apply force
+ apply (subst if_P, assumption)
+ apply (rule the1_equality)
+ apply (rule ex_ex1I)
+ apply (rule multiset_prime_factorization_exists, assumption)
+ apply (rule multiset_prime_factorization_unique)
+ apply force
+ apply force
+ apply force
+ unfolding set_of_def count_def msetprod_def
+ apply (subgoal_tac "f : multiset")
+ apply (auto simp only: Abs_multiset_inverse)
+ unfolding multiset_def apply force
+done
+
+lemma nat_prime_factors_characterization: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
+ finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+ prime_factors n = S"
+ by (rule nat_prime_factorization_unique [THEN conjunct1, symmetric],
+ assumption+)
+
+lemma nat_prime_factors_characterization':
+ "finite {p. 0 < f (p::nat)} \<Longrightarrow>
+ (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
+ prime_factors (PROD p | 0 < f p . p ^ f p) = {p. 0 < f p}"
+ apply (rule nat_prime_factors_characterization)
+ apply auto
+done
+
+(* A minor glitch:*)
+
+thm nat_prime_factors_characterization'
+ [where f = "%x. f (int (x::nat))",
+ transferred direction: nat "op <= (0::int)", rule_format]
+
+(*
+ Transfer isn't smart enough to know that the "0 < f p" should
+ remain a comparison between nats. But the transfer still works.
+*)
+
+lemma int_primes_characterization' [rule_format]:
+ "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
+ (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
+ prime_factors (PROD p | p >=0 & 0 < f p . p ^ f p) =
+ {p. p >= 0 & 0 < f p}"
+
+ apply (insert nat_prime_factors_characterization'
+ [where f = "%x. f (int (x::nat))",
+ transferred direction: nat "op <= (0::int)"])
+ apply auto
+done
+
+lemma int_prime_factors_characterization: "S = {p. 0 < f (p::int)} \<Longrightarrow>
+ finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+ prime_factors n = S"
+ apply simp
+ apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
+ apply (simp only:)
+ apply (subst int_primes_characterization')
+ apply auto
+ apply (auto simp add: int_prime_ge_0)
+done
+
+lemma nat_multiplicity_characterization: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
+ finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+ multiplicity p n = f p"
+ by (frule nat_prime_factorization_unique [THEN conjunct2, rule_format,
+ symmetric], auto)
+
+lemma nat_multiplicity_characterization': "finite {p. 0 < f (p::nat)} \<longrightarrow>
+ (ALL p. 0 < f p \<longrightarrow> prime p) \<longrightarrow>
+ multiplicity p (PROD p | 0 < f p . p ^ f p) = f p"
+ apply (rule impI)+
+ apply (rule nat_multiplicity_characterization)
+ apply auto
+done
+
+lemma int_multiplicity_characterization' [rule_format]:
+ "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
+ (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> p >= 0 \<Longrightarrow>
+ multiplicity p (PROD p | p >= 0 & 0 < f p . p ^ f p) = f p"
+
+ apply (insert nat_multiplicity_characterization'
+ [where f = "%x. f (int (x::nat))",
+ transferred direction: nat "op <= (0::int)", rule_format])
+ apply auto
+done
+
+lemma int_multiplicity_characterization: "S = {p. 0 < f (p::int)} \<Longrightarrow>
+ finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+ p >= 0 \<Longrightarrow> multiplicity p n = f p"
+ apply simp
+ apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
+ apply (simp only:)
+ apply (subst int_multiplicity_characterization')
+ apply auto
+ apply (auto simp add: int_prime_ge_0)
+done
+
+lemma nat_multiplicity_zero [simp]: "multiplicity (p::nat) 0 = 0"
+ by (simp add: multiplicity_nat_def multiset_prime_factorization_def)
+
+lemma int_multiplicity_zero [simp]: "multiplicity (p::int) 0 = 0"
+ by (simp add: multiplicity_int_def)
+
+lemma nat_multiplicity_one [simp]: "multiplicity p (1::nat) = 0"
+ by (subst nat_multiplicity_characterization [where f = "%x. 0"], auto)
+
+lemma int_multiplicity_one [simp]: "multiplicity p (1::int) = 0"
+ by (simp add: multiplicity_int_def)
+
+lemma nat_multiplicity_prime [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
+ apply (subst nat_multiplicity_characterization
+ [where f = "(%q. if q = p then 1 else 0)"])
+ apply auto
+ apply (case_tac "x = p")
+ apply auto
+done
+
+lemma int_multiplicity_prime [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
+ unfolding prime_int_def multiplicity_int_def by auto
+
+lemma nat_multiplicity_prime_power [simp]: "prime (p::nat) \<Longrightarrow>
+ multiplicity p (p^n) = n"
+ apply (case_tac "n = 0")
+ apply auto
+ apply (subst nat_multiplicity_characterization
+ [where f = "(%q. if q = p then n else 0)"])
+ apply auto
+ apply (case_tac "x = p")
+ apply auto
+done
+
+lemma int_multiplicity_prime_power [simp]: "prime (p::int) \<Longrightarrow>
+ multiplicity p (p^n) = n"
+ apply (frule int_prime_ge_0)
+ apply (auto simp add: prime_int_def multiplicity_int_def nat_power_eq)
+done
+
+lemma nat_multiplicity_nonprime [simp]: "~ prime (p::nat) \<Longrightarrow>
+ multiplicity p n = 0"
+ apply (case_tac "n = 0")
+ apply auto
+ apply (frule multiset_prime_factorization)
+ apply (auto simp add: set_of_def multiplicity_nat_def)
+done
+
+lemma int_multiplicity_nonprime [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
+ by (unfold multiplicity_int_def prime_int_def, auto)
+
+lemma nat_multiplicity_not_factor [simp]:
+ "p ~: prime_factors (n::nat) \<Longrightarrow> multiplicity p n = 0"
+ by (subst (asm) nat_prime_factors_altdef, auto)
+
+lemma int_multiplicity_not_factor [simp]:
+ "p >= 0 \<Longrightarrow> p ~: prime_factors (n::int) \<Longrightarrow> multiplicity p n = 0"
+ by (subst (asm) int_prime_factors_altdef, auto)
+
+lemma nat_multiplicity_product_aux: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
+ (prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
+ (ALL p. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
+ apply (rule nat_prime_factorization_unique)
+ apply (simp only: nat_prime_factors_altdef)
+ apply auto
+ apply (subst power_add)
+ apply (subst setprod_timesf)
+ apply (rule arg_cong2)back back
+ apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors k Un
+ (prime_factors l - prime_factors k)")
+ apply (erule ssubst)
+ apply (subst setprod_Un_disjoint)
+ apply auto
+ apply (subgoal_tac "(\<Prod>p\<in>prime_factors l - prime_factors k. p ^ multiplicity p k) =
+ (\<Prod>p\<in>prime_factors l - prime_factors k. 1)")
+ apply (erule ssubst)
+ apply (simp add: setprod_1)
+ apply (erule nat_prime_factorization)
+ apply (rule setprod_cong, auto)
+ apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un
+ (prime_factors k - prime_factors l)")
+ apply (erule ssubst)
+ apply (subst setprod_Un_disjoint)
+ apply auto
+ apply (subgoal_tac "(\<Prod>p\<in>prime_factors k - prime_factors l. p ^ multiplicity p l) =
+ (\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
+ apply (erule ssubst)
+ apply (simp add: setprod_1)
+ apply (erule nat_prime_factorization)
+ apply (rule setprod_cong, auto)
+done
+
+(* transfer doesn't have the same problem here with the right
+ choice of rules. *)
+
+lemma int_multiplicity_product_aux:
+ assumes "(k::int) > 0" and "l > 0"
+ shows
+ "(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
+ (ALL p >= 0. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
+
+ apply (rule nat_multiplicity_product_aux [transferred, of l k])
+ using prems apply auto
+done
+
+lemma nat_prime_factors_product: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
+ prime_factors k Un prime_factors l"
+ by (rule nat_multiplicity_product_aux [THEN conjunct1, symmetric])
+
+lemma int_prime_factors_product: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
+ prime_factors k Un prime_factors l"
+ by (rule int_multiplicity_product_aux [THEN conjunct1, symmetric])
+
+lemma nat_multiplicity_product: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) =
+ multiplicity p k + multiplicity p l"
+ by (rule nat_multiplicity_product_aux [THEN conjunct2, rule_format,
+ symmetric])
+
+lemma int_multiplicity_product: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow>
+ multiplicity p (k * l) = multiplicity p k + multiplicity p l"
+ by (rule int_multiplicity_product_aux [THEN conjunct2, rule_format,
+ symmetric])
+
+lemma nat_multiplicity_setprod: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow>
+ multiplicity (p::nat) (PROD x : S. f x) =
+ (SUM x : S. multiplicity p (f x))"
+ apply (induct set: finite)
+ apply auto
+ apply (subst nat_multiplicity_product)
+ apply auto
+done
+
+(* Transfer is delicate here for two reasons: first, because there is
+ an implicit quantifier over functions (f), and, second, because the
+ product over the multiplicity should not be translated to an integer
+ product.
+
+ The way to handle the first is to use quantifier rules for functions.
+ The way to handle the second is to turn off the offending rule.
+*)
+
+lemma transfer_nat_int_sum_prod_closure3:
+ "(SUM x : A. int (f x)) >= 0"
+ "(PROD x : A. int (f x)) >= 0"
+ apply (rule setsum_nonneg, auto)
+ apply (rule setprod_nonneg, auto)
+done
+
+declare TransferMorphism_nat_int[transfer
+ add return: transfer_nat_int_sum_prod_closure3
+ del: transfer_nat_int_sum_prod2 (1)]
+
+lemma int_multiplicity_setprod: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow>
+ (ALL x : S. f x > 0) \<Longrightarrow>
+ multiplicity (p::int) (PROD x : S. f x) =
+ (SUM x : S. multiplicity p (f x))"
+
+ apply (frule nat_multiplicity_setprod
+ [where f = "%x. nat(int(nat(f x)))",
+ transferred direction: nat "op <= (0::int)"])
+ apply auto
+ apply (subst (asm) setprod_cong)
+ apply (rule refl)
+ apply (rule if_P)
+ apply auto
+ apply (rule setsum_cong)
+ apply auto
+done
+
+declare TransferMorphism_nat_int[transfer
+ add return: transfer_nat_int_sum_prod2 (1)]
+
+lemma nat_multiplicity_prod_prime_powers:
+ "finite S \<Longrightarrow> (ALL p : S. prime (p::nat)) \<Longrightarrow>
+ multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
+ apply (subgoal_tac "(PROD p : S. p ^ f p) =
+ (PROD p : S. p ^ (%x. if x : S then f x else 0) p)")
+ apply (erule ssubst)
+ apply (subst nat_multiplicity_characterization)
+ prefer 5 apply (rule refl)
+ apply (rule refl)
+ apply auto
+ apply (subst setprod_mono_one_right)
+ apply assumption
+ prefer 3
+ apply (rule setprod_cong)
+ apply (rule refl)
+ apply auto
+done
+
+(* Here the issue with transfer is the implicit quantifier over S *)
+
+lemma int_multiplicity_prod_prime_powers:
+ "(p::int) >= 0 \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
+ multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
+
+ apply (subgoal_tac "int ` nat ` S = S")
+ apply (frule nat_multiplicity_prod_prime_powers [where f = "%x. f(int x)"
+ and S = "nat ` S", transferred])
+ apply auto
+ apply (subst prime_int_def [symmetric])
+ apply auto
+ apply (subgoal_tac "xb >= 0")
+ apply force
+ apply (rule int_prime_ge_0)
+ apply force
+ apply (subst transfer_nat_int_set_return_embed)
+ apply (unfold nat_set_def, auto)
+done
+
+lemma nat_multiplicity_distinct_prime_power: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
+ p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
+ apply (subgoal_tac "q^n = setprod (%x. x^n) {q}")
+ apply (erule ssubst)
+ apply (subst nat_multiplicity_prod_prime_powers)
+ apply auto
+done
+
+lemma int_multiplicity_distinct_prime_power: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
+ p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
+ apply (frule int_prime_ge_0 [of q])
+ apply (frule nat_multiplicity_distinct_prime_power [transferred leaving: n])
+ prefer 4
+ apply assumption
+ apply auto
+done
+
+lemma nat_dvd_multiplicity:
+ "(0::nat) < y \<Longrightarrow> x dvd y \<Longrightarrow> multiplicity p x <= multiplicity p y"
+ apply (case_tac "x = 0")
+ apply (auto simp add: dvd_def nat_multiplicity_product)
+done
+
+lemma int_dvd_multiplicity:
+ "(0::int) < y \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> p >= 0 \<Longrightarrow>
+ multiplicity p x <= multiplicity p y"
+ apply (case_tac "x = 0")
+ apply (auto simp add: dvd_def)
+ apply (subgoal_tac "0 < k")
+ apply (auto simp add: int_multiplicity_product)
+ apply (erule zero_less_mult_pos)
+ apply arith
+done
+
+lemma nat_dvd_prime_factors [intro]:
+ "0 < (y::nat) \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
+ apply (simp only: nat_prime_factors_altdef)
+ apply auto
+ apply (frule nat_dvd_multiplicity)
+ apply auto
+(* It is a shame that auto and arith don't get this. *)
+ apply (erule order_less_le_trans)back
+ apply assumption
+done
+
+lemma int_dvd_prime_factors [intro]:
+ "0 < (y::int) \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
+ apply (auto simp add: int_prime_factors_altdef)
+ apply (erule order_less_le_trans)
+ apply (rule int_dvd_multiplicity)
+ apply auto
+done
+
+lemma nat_multiplicity_dvd: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
+ ALL p. multiplicity p x <= multiplicity p y \<Longrightarrow>
+ x dvd y"
+ apply (subst nat_prime_factorization [of x], assumption)
+ apply (subst nat_prime_factorization [of y], assumption)
+ apply (rule setprod_dvd_setprod_subset2)
+ apply force
+ apply (subst nat_prime_factors_altdef)+
+ apply auto
+(* Again, a shame that auto and arith don't get this. *)
+ apply (drule_tac x = xa in spec, auto)
+ apply (rule le_imp_power_dvd)
+ apply blast
+done
+
+lemma int_multiplicity_dvd: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
+ ALL p >= 0. multiplicity p x <= multiplicity p y \<Longrightarrow>
+ x dvd y"
+ apply (subst int_prime_factorization [of x], assumption)
+ apply (subst int_prime_factorization [of y], assumption)
+ apply (rule setprod_dvd_setprod_subset2)
+ apply force
+ apply (subst int_prime_factors_altdef)+
+ apply auto
+ apply (rule dvd_power_le)
+ apply auto
+ apply (drule_tac x = xa in spec)
+ apply (erule impE)
+ apply auto
+done
+
+lemma nat_multiplicity_dvd': "(0::nat) < x \<Longrightarrow>
+ \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
+ apply (cases "y = 0")
+ apply auto
+ apply (rule nat_multiplicity_dvd, auto)
+ apply (case_tac "prime p")
+ apply auto
+done
+
+lemma int_multiplicity_dvd': "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
+ \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
+ apply (cases "y = 0")
+ apply auto
+ apply (rule int_multiplicity_dvd, auto)
+ apply (case_tac "prime p")
+ apply auto
+done
+
+lemma nat_dvd_multiplicity_eq: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
+ (x dvd y) = (ALL p. multiplicity p x <= multiplicity p y)"
+ by (auto intro: nat_dvd_multiplicity nat_multiplicity_dvd)
+
+lemma int_dvd_multiplicity_eq: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
+ (x dvd y) = (ALL p >= 0. multiplicity p x <= multiplicity p y)"
+ by (auto intro: int_dvd_multiplicity int_multiplicity_dvd)
+
+lemma nat_prime_factors_altdef2: "(n::nat) > 0 \<Longrightarrow>
+ (p : prime_factors n) = (prime p & p dvd n)"
+ apply (case_tac "prime p")
+ apply auto
+ apply (subst nat_prime_factorization [where n = n], assumption)
+ apply (rule dvd_trans)
+ apply (rule dvd_power [where x = p and n = "multiplicity p n"])
+ apply (subst (asm) nat_prime_factors_altdef, force)
+ apply (rule dvd_setprod)
+ apply auto
+ apply (subst nat_prime_factors_altdef)
+ apply (subst (asm) nat_dvd_multiplicity_eq)
+ apply auto
+ apply (drule spec [where x = p])
+ apply auto
+done
+
+lemma int_prime_factors_altdef2:
+ assumes "(n::int) > 0"
+ shows "(p : prime_factors n) = (prime p & p dvd n)"
+
+ apply (case_tac "p >= 0")
+ apply (rule nat_prime_factors_altdef2 [transferred])
+ using prems apply auto
+ apply (auto simp add: int_prime_ge_0 int_prime_factors_ge_0)
+done
+
+lemma nat_multiplicity_eq:
+ fixes x and y::nat
+ assumes [arith]: "x > 0" "y > 0" and
+ mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+ shows "x = y"
+
+ apply (rule dvd_anti_sym)
+ apply (auto intro: nat_multiplicity_dvd')
+done
+
+lemma int_multiplicity_eq:
+ fixes x and y::int
+ assumes [arith]: "x > 0" "y > 0" and
+ mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+ shows "x = y"
+
+ apply (rule dvd_anti_sym [transferred])
+ apply (auto intro: int_multiplicity_dvd')
+done
+
+
+subsection {* An application *}
+
+lemma nat_gcd_eq:
+ assumes pos [arith]: "x > 0" "y > 0"
+ shows "gcd (x::nat) y =
+ (PROD p: prime_factors x Un prime_factors y.
+ p ^ (min (multiplicity p x) (multiplicity p y)))"
+proof -
+ def z == "(PROD p: prime_factors (x::nat) Un prime_factors y.
+ p ^ (min (multiplicity p x) (multiplicity p y)))"
+ have [arith]: "z > 0"
+ unfolding z_def by (rule setprod_pos_nat, auto)
+ have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
+ min (multiplicity p x) (multiplicity p y)"
+ unfolding z_def
+ apply (subst nat_multiplicity_prod_prime_powers)
+ apply (auto simp add: nat_multiplicity_not_factor)
+ done
+ have "z dvd x"
+ by (intro nat_multiplicity_dvd', auto simp add: aux)
+ moreover have "z dvd y"
+ by (intro nat_multiplicity_dvd', auto simp add: aux)
+ moreover have "ALL w. w dvd x & w dvd y \<longrightarrow> w dvd z"
+ apply auto
+ apply (case_tac "w = 0", auto)
+ apply (erule nat_multiplicity_dvd')
+ apply (auto intro: nat_dvd_multiplicity simp add: aux)
+ done
+ ultimately have "z = gcd x y"
+ by (subst nat_gcd_unique [symmetric], blast)
+ thus ?thesis
+ unfolding z_def by auto
+qed
+
+lemma nat_lcm_eq:
+ assumes pos [arith]: "x > 0" "y > 0"
+ shows "lcm (x::nat) y =
+ (PROD p: prime_factors x Un prime_factors y.
+ p ^ (max (multiplicity p x) (multiplicity p y)))"
+proof -
+ def z == "(PROD p: prime_factors (x::nat) Un prime_factors y.
+ p ^ (max (multiplicity p x) (multiplicity p y)))"
+ have [arith]: "z > 0"
+ unfolding z_def by (rule setprod_pos_nat, auto)
+ have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
+ max (multiplicity p x) (multiplicity p y)"
+ unfolding z_def
+ apply (subst nat_multiplicity_prod_prime_powers)
+ apply (auto simp add: nat_multiplicity_not_factor)
+ done
+ have "x dvd z"
+ by (intro nat_multiplicity_dvd', auto simp add: aux)
+ moreover have "y dvd z"
+ by (intro nat_multiplicity_dvd', auto simp add: aux)
+ moreover have "ALL w. x dvd w & y dvd w \<longrightarrow> z dvd w"
+ apply auto
+ apply (case_tac "w = 0", auto)
+ apply (rule nat_multiplicity_dvd')
+ apply (auto intro: nat_dvd_multiplicity simp add: aux)
+ done
+ ultimately have "z = lcm x y"
+ by (subst nat_lcm_unique [symmetric], blast)
+ thus ?thesis
+ unfolding z_def by auto
+qed
+
+lemma nat_multiplicity_gcd:
+ assumes [arith]: "x > 0" "y > 0"
+ shows "multiplicity (p::nat) (gcd x y) =
+ min (multiplicity p x) (multiplicity p y)"
+
+ apply (subst nat_gcd_eq)
+ apply auto
+ apply (subst nat_multiplicity_prod_prime_powers)
+ apply auto
+done
+
+lemma nat_multiplicity_lcm:
+ assumes [arith]: "x > 0" "y > 0"
+ shows "multiplicity (p::nat) (lcm x y) =
+ max (multiplicity p x) (multiplicity p y)"
+
+ apply (subst nat_lcm_eq)
+ apply auto
+ apply (subst nat_multiplicity_prod_prime_powers)
+ apply auto
+done
+
+lemma nat_gcd_lcm_distrib: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
+ apply (case_tac "x = 0 | y = 0 | z = 0")
+ apply auto
+ apply (rule nat_multiplicity_eq)
+ apply (auto simp add: nat_multiplicity_gcd nat_multiplicity_lcm
+ nat_lcm_pos)
+done
+
+lemma int_gcd_lcm_distrib: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
+ apply (subst (1 2 3) int_gcd_abs)
+ apply (subst int_lcm_abs)
+ apply (subst (2) abs_of_nonneg)
+ apply force
+ apply (rule nat_gcd_lcm_distrib [transferred])
+ apply auto
+done
+
+end
--- a/src/HOL/Nominal/Nominal.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nominal/Nominal.thy Tue Jun 23 15:48:55 2009 +0100
@@ -3,7 +3,7 @@
uses
("nominal_thmdecls.ML")
("nominal_atoms.ML")
- ("nominal_package.ML")
+ ("nominal.ML")
("nominal_induct.ML")
("nominal_permeq.ML")
("nominal_fresh_fun.ML")
@@ -3670,7 +3670,7 @@
lemma allE_Nil: assumes "\<forall>x. P x" obtains "P []"
using assms ..
-use "nominal_package.ML"
+use "nominal.ML"
(******************************************************)
(* primitive recursive functions on nominal datatypes *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Nominal/nominal.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,2095 @@
+(* Title: HOL/Nominal/nominal.ML
+ Author: Stefan Berghofer and Christian Urban, TU Muenchen
+
+Nominal datatype package for Isabelle/HOL.
+*)
+
+signature NOMINAL =
+sig
+ val add_nominal_datatype : Datatype.config -> string list ->
+ (string list * bstring * mixfix *
+ (bstring * string list * mixfix) list) list -> theory -> theory
+ type descr
+ type nominal_datatype_info
+ val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
+ val get_nominal_datatype : theory -> string -> nominal_datatype_info option
+ val mk_perm: typ list -> term -> term -> term
+ val perm_of_pair: term * term -> term
+ val mk_not_sym: thm list -> thm list
+ val perm_simproc: simproc
+ val fresh_const: typ -> typ -> term
+ val fresh_star_const: typ -> typ -> term
+end
+
+structure Nominal : NOMINAL =
+struct
+
+val finite_emptyI = thm "finite.emptyI";
+val finite_Diff = thm "finite_Diff";
+val finite_Un = thm "finite_Un";
+val Un_iff = thm "Un_iff";
+val In0_eq = thm "In0_eq";
+val In1_eq = thm "In1_eq";
+val In0_not_In1 = thm "In0_not_In1";
+val In1_not_In0 = thm "In1_not_In0";
+val Un_assoc = thm "Un_assoc";
+val Collect_disj_eq = thm "Collect_disj_eq";
+val empty_def = thm "empty_def";
+val empty_iff = thm "empty_iff";
+
+open DatatypeAux;
+open NominalAtoms;
+
+(** FIXME: Datatype should export this function **)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
+
+fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (RuleCases.case_names o exhaust_cases descr o #1)
+ (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
+
+end;
+
+(* theory data *)
+
+type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
+
+type nominal_datatype_info =
+ {index : int,
+ descr : descr,
+ sorts : (string * sort) list,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ induction : thm,
+ distinct : thm list,
+ inject : thm list};
+
+structure NominalDatatypesData = TheoryDataFun
+(
+ type T = nominal_datatype_info Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ tabs : T = Symtab.merge (K true) tabs;
+);
+
+val get_nominal_datatypes = NominalDatatypesData.get;
+val put_nominal_datatypes = NominalDatatypesData.put;
+val map_nominal_datatypes = NominalDatatypesData.map;
+val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
+
+
+(**** make datatype info ****)
+
+fun make_dt_info descr sorts induct reccomb_names rec_thms
+ (((i, (_, (tname, _, _))), distinct), inject) =
+ (tname,
+ {index = i,
+ descr = descr,
+ sorts = sorts,
+ rec_names = reccomb_names,
+ rec_rewrites = rec_thms,
+ induction = induct,
+ distinct = distinct,
+ inject = inject});
+
+(*******************************)
+
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+
+
+(** simplification procedure for sorting permutations **)
+
+val dj_cp = thm "dj_cp";
+
+fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
+ Type ("fun", [_, U])])) = (T, U);
+
+fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
+ | permTs_of _ = [];
+
+fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
+ let
+ val (aT as Type (a, []), S) = dest_permT T;
+ val (bT as Type (b, []), _) = dest_permT U
+ in if aT mem permTs_of u andalso aT <> bT then
+ let
+ val cp = cp_inst_of thy a b;
+ val dj = dj_thm_of thy b a;
+ val dj_cp' = [cp, dj] MRS dj_cp;
+ val cert = SOME o cterm_of thy
+ in
+ SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
+ [cert t, cert r, cert s] dj_cp'))
+ end
+ else NONE
+ end
+ | perm_simproc' thy ss _ = NONE;
+
+val perm_simproc =
+ Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
+
+val meta_spec = thm "meta_spec";
+
+fun projections rule =
+ ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
+ |> map (standard #> RuleCases.save rule);
+
+val supp_prod = thm "supp_prod";
+val fresh_prod = thm "fresh_prod";
+val supports_fresh = thm "supports_fresh";
+val supports_def = thm "Nominal.supports_def";
+val fresh_def = thm "fresh_def";
+val supp_def = thm "supp_def";
+val rev_simps = thms "rev.simps";
+val app_simps = thms "append.simps";
+val at_fin_set_supp = thm "at_fin_set_supp";
+val at_fin_set_fresh = thm "at_fin_set_fresh";
+val abs_fun_eq1 = thm "abs_fun_eq1";
+
+val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
+
+fun mk_perm Ts t u =
+ let
+ val T = fastype_of1 (Ts, t);
+ val U = fastype_of1 (Ts, u)
+ in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
+
+fun perm_of_pair (x, y) =
+ let
+ val T = fastype_of x;
+ val pT = mk_permT T
+ in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
+ HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
+ end;
+
+fun mk_not_sym ths = maps (fn th => case prop_of th of
+ _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
+ | _ => [th]) ths;
+
+fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
+fun fresh_star_const T U =
+ Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT);
+
+fun gen_add_nominal_datatype prep_typ config new_type_names dts thy =
+ let
+ (* this theory is used just for parsing *)
+
+ val tmp_thy = thy |>
+ Theory.copy |>
+ Sign.add_types (map (fn (tvs, tname, mx, _) =>
+ (Binding.name tname, length tvs, mx)) dts);
+
+ val atoms = atoms_of thy;
+
+ fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
+ let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
+ in (constrs @ [(cname, cargs', mx)], sorts') end
+
+ fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
+ let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
+ in (dts @ [(tvs, tname, mx, constrs')], sorts') end
+
+ val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
+ val tyvars = map (map (fn s =>
+ (s, the (AList.lookup (op =) sorts s))) o #1) dts';
+
+ fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
+ fun augment_sort_typ thy S =
+ let val S = Sign.certify_sort thy S
+ in map_type_tfree (fn (s, S') => TFree (s,
+ if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
+ end;
+ fun augment_sort thy S = map_types (augment_sort_typ thy S);
+
+ val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
+ val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
+ map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
+
+ val ps = map (fn (_, n, _, _) =>
+ (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
+ val rps = map Library.swap ps;
+
+ fun replace_types (Type ("Nominal.ABS", [T, U])) =
+ Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
+ | replace_types (Type (s, Ts)) =
+ Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
+ | replace_types T = T;
+
+ val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
+ map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"),
+ map replace_types cargs, NoSyn)) constrs)) dts';
+
+ val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
+ val full_new_type_names' = map (Sign.full_bname thy) new_type_names';
+
+ val ({induction, ...},thy1) =
+ Datatype.add_datatype config new_type_names' dts'' thy;
+
+ val SOME {descr, ...} = Symtab.lookup
+ (Datatype.get_datatypes thy1) (hd full_new_type_names');
+ fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
+
+ val big_name = space_implode "_" new_type_names;
+
+
+ (**** define permutation functions ****)
+
+ val permT = mk_permT (TFree ("'x", HOLogic.typeS));
+ val pi = Free ("pi", permT);
+ val perm_types = map (fn (i, _) =>
+ let val T = nth_dtyp i
+ in permT --> T --> T end) descr;
+ val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
+ "perm_" ^ name_of_typ (nth_dtyp i)) descr);
+ val perm_names = replicate (length new_type_names) "Nominal.perm" @
+ map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
+ val perm_names_types = perm_names ~~ perm_types;
+ val perm_names_types' = perm_names' ~~ perm_types;
+
+ val perm_eqs = maps (fn (i, (_, _, constrs)) =>
+ let val T = nth_dtyp i
+ in map (fn (cname, dts) =>
+ let
+ val Ts = map (typ_of_dtyp descr sorts) dts;
+ val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
+ val args = map Free (names ~~ Ts);
+ val c = Const (cname, Ts ---> T);
+ fun perm_arg (dt, x) =
+ let val T = type_of x
+ in if is_rec_type dt then
+ let val (Us, _) = strip_type T
+ in list_abs (map (pair "x") Us,
+ Free (nth perm_names_types' (body_index dt)) $ pi $
+ list_comb (x, map (fn (i, U) =>
+ Const ("Nominal.perm", permT --> U --> U) $
+ (Const ("List.rev", permT --> permT) $ pi) $
+ Bound i) ((length Us - 1 downto 0) ~~ Us)))
+ end
+ else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
+ end;
+ in
+ (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Free (nth perm_names_types' i) $
+ Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
+ list_comb (c, args),
+ list_comb (c, map perm_arg (dts ~~ args)))))
+ end) constrs
+ end) descr;
+
+ val (perm_simps, thy2) =
+ Primrec.add_primrec_overloaded
+ (map (fn (s, sT) => (s, sT, false))
+ (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
+ (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
+
+ (**** prove that permutation functions introduced by unfolding are ****)
+ (**** equivalent to already existing permutation functions ****)
+
+ val _ = warning ("length descr: " ^ string_of_int (length descr));
+ val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
+
+ val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
+ val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def";
+
+ val unfolded_perm_eq_thms =
+ if length descr = length new_type_names then []
+ else map standard (List.drop (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (c as (s, T), x) =>
+ let val [T1, T2] = binder_types T
+ in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
+ Const ("Nominal.perm", T) $ pi $ Free (x, T2))
+ end)
+ (perm_names_types ~~ perm_indnames))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac
+ (simpset_of thy2 addsimps [perm_fun_def]))])),
+ length new_type_names));
+
+ (**** prove [] \<bullet> t = t ****)
+
+ val _ = warning "perm_empty_thms";
+
+ val perm_empty_thms = List.concat (map (fn a =>
+ let val permT = mk_permT (Type (a, []))
+ in map standard (List.take (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a]
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) => HOLogic.mk_eq
+ (Const (s, permT --> T --> T) $
+ Const ("List.list.Nil", permT) $ Free (x, T),
+ Free (x, T)))
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
+ length new_type_names))
+ end)
+ atoms);
+
+ (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
+
+ val _ = warning "perm_append_thms";
+
+ (*FIXME: these should be looked up statically*)
+ val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst";
+ val pt2 = PureThy.get_thm thy2 "pt2";
+
+ val perm_append_thms = List.concat (map (fn a =>
+ let
+ val permT = mk_permT (Type (a, []));
+ val pi1 = Free ("pi1", permT);
+ val pi2 = Free ("pi2", permT);
+ val pt_inst = pt_inst_of thy2 a;
+ val pt2' = pt_inst RS pt2;
+ val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
+ in List.take (map standard (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a]
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let val perm = Const (s, permT --> T --> T)
+ in HOLogic.mk_eq
+ (perm $ (Const ("List.append", permT --> permT --> permT) $
+ pi1 $ pi2) $ Free (x, T),
+ perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
+ end)
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
+ length new_type_names)
+ end) atoms);
+
+ (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
+
+ val _ = warning "perm_eq_thms";
+
+ val pt3 = PureThy.get_thm thy2 "pt3";
+ val pt3_rev = PureThy.get_thm thy2 "pt3_rev";
+
+ val perm_eq_thms = List.concat (map (fn a =>
+ let
+ val permT = mk_permT (Type (a, []));
+ val pi1 = Free ("pi1", permT);
+ val pi2 = Free ("pi2", permT);
+ val at_inst = at_inst_of thy2 a;
+ val pt_inst = pt_inst_of thy2 a;
+ val pt3' = pt_inst RS pt3;
+ val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
+ val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
+ in List.take (map standard (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
+ (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
+ permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
+ HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let val perm = Const (s, permT --> T --> T)
+ in HOLogic.mk_eq
+ (perm $ pi1 $ Free (x, T),
+ perm $ pi2 $ Free (x, T))
+ end)
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames))))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
+ length new_type_names)
+ end) atoms);
+
+ (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
+
+ val cp1 = PureThy.get_thm thy2 "cp1";
+ val dj_cp = PureThy.get_thm thy2 "dj_cp";
+ val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose";
+ val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
+ val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
+
+ fun composition_instance name1 name2 thy =
+ let
+ val cp_class = cp_class_of thy name1 name2;
+ val pt_class =
+ if name1 = name2 then [pt_class_of thy name1]
+ else [];
+ val permT1 = mk_permT (Type (name1, []));
+ val permT2 = mk_permT (Type (name2, []));
+ val Ts = map body_type perm_types;
+ val cp_inst = cp_inst_of thy name1 name2;
+ val simps = simpset_of thy addsimps (perm_fun_def ::
+ (if name1 <> name2 then
+ let val dj = dj_thm_of thy name2 name1
+ in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
+ else
+ let
+ val at_inst = at_inst_of thy name1;
+ val pt_inst = pt_inst_of thy name1;
+ in
+ [cp_inst RS cp1 RS sym,
+ at_inst RS (pt_inst RS pt_perm_compose) RS sym,
+ at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
+ end))
+ val sort = Sign.certify_sort thy (cp_class :: pt_class);
+ val thms = split_conj_thm (Goal.prove_global thy [] []
+ (augment_sort thy sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let
+ val pi1 = Free ("pi1", permT1);
+ val pi2 = Free ("pi2", permT2);
+ val perm1 = Const (s, permT1 --> T --> T);
+ val perm2 = Const (s, permT2 --> T --> T);
+ val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
+ in HOLogic.mk_eq
+ (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
+ perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
+ end)
+ (perm_names ~~ Ts ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac simps)]))
+ in
+ fold (fn (s, tvs) => fn thy => AxClass.prove_arity
+ (s, map (inter_sort thy sort o snd) tvs, [cp_class])
+ (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
+ (full_new_type_names' ~~ tyvars) thy
+ end;
+
+ val (perm_thmss,thy3) = thy2 |>
+ fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
+ fold (fn atom => fn thy =>
+ let val pt_name = pt_class_of thy atom
+ in
+ fold (fn (s, tvs) => fn thy => AxClass.prove_arity
+ (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
+ (EVERY
+ [Class.intro_classes_tac [],
+ resolve_tac perm_empty_thms 1,
+ resolve_tac perm_append_thms 1,
+ resolve_tac perm_eq_thms 1, assume_tac 1]) thy)
+ (full_new_type_names' ~~ tyvars) thy
+ end) atoms |>
+ PureThy.add_thmss
+ [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
+ unfolded_perm_eq_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
+ perm_empty_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
+ perm_append_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
+ perm_eq_thms), [Simplifier.simp_add])];
+
+ (**** Define representing sets ****)
+
+ val _ = warning "representing sets";
+
+ val rep_set_names = DatatypeProp.indexify_names
+ (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
+ val big_rep_name =
+ space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
+ (fn (i, ("Nominal.noption", _, _)) => NONE
+ | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
+ val _ = warning ("big_rep_name: " ^ big_rep_name);
+
+ fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
+ (case AList.lookup op = descr i of
+ SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
+ apfst (cons dt) (strip_option dt')
+ | _ => ([], dtf))
+ | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
+ apfst (cons dt) (strip_option dt')
+ | strip_option dt = ([], dt);
+
+ val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
+ (List.concat (map (fn (_, (_, _, cs)) => List.concat
+ (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
+ val dt_atoms = map (fst o dest_Type) dt_atomTs;
+
+ fun make_intr s T (cname, cargs) =
+ let
+ fun mk_prem (dt, (j, j', prems, ts)) =
+ let
+ val (dts, dt') = strip_option dt;
+ val (dts', dt'') = strip_dtyp dt';
+ val Ts = map (typ_of_dtyp descr sorts) dts;
+ val Us = map (typ_of_dtyp descr sorts) dts';
+ val T = typ_of_dtyp descr sorts dt'';
+ val free = mk_Free "x" (Us ---> T) j;
+ val free' = app_bnds free (length Us);
+ fun mk_abs_fun (T, (i, t)) =
+ let val U = fastype_of t
+ in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
+ Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
+ end
+ in (j + 1, j' + length Ts,
+ case dt'' of
+ DtRec k => list_all (map (pair "x") Us,
+ HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
+ T --> HOLogic.boolT) $ free')) :: prems
+ | _ => prems,
+ snd (List.foldr mk_abs_fun (j', free) Ts) :: ts)
+ end;
+
+ val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs;
+ val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
+ list_comb (Const (cname, map fastype_of ts ---> T), ts))
+ in Logic.list_implies (prems, concl)
+ end;
+
+ val (intr_ts, (rep_set_names', recTs')) =
+ apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
+ (fn ((_, ("Nominal.noption", _, _)), _) => NONE
+ | ((i, (_, _, constrs)), rep_set_name) =>
+ let val T = nth_dtyp i
+ in SOME (map (make_intr rep_set_name T) constrs,
+ (rep_set_name, T))
+ end)
+ (descr ~~ rep_set_names))));
+ val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
+
+ val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = false, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
+ (rep_set_names' ~~ recTs'))
+ [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
+
+ (**** Prove that representing set is closed under permutation ****)
+
+ val _ = warning "proving closure under permutation...";
+
+ val abs_perm = PureThy.get_thms thy4 "abs_perm";
+
+ val perm_indnames' = List.mapPartial
+ (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
+ (perm_indnames ~~ descr);
+
+ fun mk_perm_closed name = map (fn th => standard (th RS mp))
+ (List.take (split_conj_thm (Goal.prove_global thy4 [] []
+ (augment_sort thy4
+ (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
+ (fn ((s, T), x) =>
+ let
+ val S = Const (s, T --> HOLogic.boolT);
+ val permT = mk_permT (Type (name, []))
+ in HOLogic.mk_imp (S $ Free (x, T),
+ S $ (Const ("Nominal.perm", permT --> T --> T) $
+ Free ("pi", permT) $ Free (x, T)))
+ end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
+ (fn _ => EVERY
+ [indtac rep_induct [] 1,
+ ALLGOALS (simp_tac (simpset_of thy4 addsimps
+ (symmetric perm_fun_def :: abs_perm))),
+ ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
+ length new_type_names));
+
+ val perm_closed_thmss = map mk_perm_closed atoms;
+
+ (**** typedef ****)
+
+ val _ = warning "defining type...";
+
+ val (typedefs, thy6) =
+ thy4
+ |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
+ Typedef.add_typedef false (SOME (Binding.name name'))
+ (Binding.name name, map fst tvs, mx)
+ (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
+ Const (cname, U --> HOLogic.boolT)) NONE
+ (rtac exI 1 THEN rtac CollectI 1 THEN
+ QUIET_BREADTH_FIRST (has_fewer_prems 1)
+ (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
+ let
+ val permT = mk_permT
+ (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS));
+ val pi = Free ("pi", permT);
+ val T = Type (Sign.intern_type thy name, map TFree tvs);
+ in apfst (pair r o hd)
+ (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
+ (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
+ Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
+ (Const ("Nominal.perm", permT --> U --> U) $ pi $
+ (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
+ Free ("x", T))))), [])] thy)
+ end))
+ (types_syntax ~~ tyvars ~~
+ List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
+ new_type_names);
+
+ val perm_defs = map snd typedefs;
+ val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
+ val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
+ val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
+
+
+ (** prove that new types are in class pt_<name> **)
+
+ val _ = warning "prove that new types are in class pt_<name> ...";
+
+ fun pt_instance (atom, perm_closed_thms) =
+ fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
+ perm_def), name), tvs), perm_closed) => fn thy =>
+ let
+ val pt_class = pt_class_of thy atom;
+ val sort = Sign.certify_sort thy
+ (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom))
+ in AxClass.prove_arity
+ (Sign.intern_type thy name,
+ map (inter_sort thy sort o snd) tvs, [pt_class])
+ (EVERY [Class.intro_classes_tac [],
+ rewrite_goals_tac [perm_def],
+ asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
+ asm_full_simp_tac (simpset_of thy addsimps
+ [Rep RS perm_closed RS Abs_inverse]) 1,
+ asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
+ ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
+ end)
+ (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
+ new_type_names ~~ tyvars ~~ perm_closed_thms);
+
+
+ (** prove that new types are in class cp_<name1>_<name2> **)
+
+ val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
+
+ fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
+ let
+ val cp_class = cp_class_of thy atom1 atom2;
+ val sort = Sign.certify_sort thy
+ (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
+ (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
+ pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2)));
+ val cp1' = cp_inst_of thy atom1 atom2 RS cp1
+ in fold (fn ((((((Abs_inverse, Rep),
+ perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
+ AxClass.prove_arity
+ (Sign.intern_type thy name,
+ map (inter_sort thy sort o snd) tvs, [cp_class])
+ (EVERY [Class.intro_classes_tac [],
+ rewrite_goals_tac [perm_def],
+ asm_full_simp_tac (simpset_of thy addsimps
+ ((Rep RS perm_closed1 RS Abs_inverse) ::
+ (if atom1 = atom2 then []
+ else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
+ cong_tac 1,
+ rtac refl 1,
+ rtac cp1' 1]) thy)
+ (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
+ tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
+ end;
+
+ val thy7 = fold (fn x => fn thy => thy |>
+ pt_instance x |>
+ fold (cp_instance x) (atoms ~~ perm_closed_thmss))
+ (atoms ~~ perm_closed_thmss) thy6;
+
+ (**** constructors ****)
+
+ fun mk_abs_fun (x, t) =
+ let
+ val T = fastype_of x;
+ val U = fastype_of t
+ in
+ Const ("Nominal.abs_fun", T --> U --> T -->
+ Type ("Nominal.noption", [U])) $ x $ t
+ end;
+
+ val (ty_idxs, _) = List.foldl
+ (fn ((i, ("Nominal.noption", _, _)), p) => p
+ | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
+
+ fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
+ | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
+ | reindex dt = dt;
+
+ fun strip_suffix i s = implode (List.take (explode s, size s - i));
+
+ (** strips the "_Rep" in type names *)
+ fun strip_nth_name i s =
+ let val xs = Long_Name.explode s;
+ in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
+
+ val (descr'', ndescr) = ListPair.unzip (map_filter
+ (fn (i, ("Nominal.noption", _, _)) => NONE
+ | (i, (s, dts, constrs)) =>
+ let
+ val SOME index = AList.lookup op = ty_idxs i;
+ val (constrs2, constrs1) =
+ map_split (fn (cname, cargs) =>
+ apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
+ (fold_map (fn dt => fn dts =>
+ let val (dts', dt') = strip_option dt
+ in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
+ cargs [])) constrs
+ in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)),
+ (index, constrs2))
+ end) descr);
+
+ val (descr1, descr2) = chop (length new_type_names) descr'';
+ val descr' = [descr1, descr2];
+
+ fun partition_cargs idxs xs = map (fn (i, j) =>
+ (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
+
+ val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
+ map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
+ (constrs ~~ idxss)))) (descr'' ~~ ndescr);
+
+ fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
+
+ val rep_names = map (fn s =>
+ Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
+ val abs_names = map (fn s =>
+ Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
+
+ val recTs = get_rec_types descr'' sorts;
+ val newTs' = Library.take (length new_type_names, recTs');
+ val newTs = Library.take (length new_type_names, recTs);
+
+ val full_new_type_names = map (Sign.full_bname thy) new_type_names;
+
+ fun make_constr_def tname T T' ((thy, defs, eqns),
+ (((cname_rep, _), (cname, cargs)), (cname', mx))) =
+ let
+ fun constr_arg ((dts, dt), (j, l_args, r_args)) =
+ let
+ val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
+ (dts ~~ (j upto j + length dts - 1))
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ x :: l_args,
+ List.foldr mk_abs_fun
+ (case dt of
+ DtRec k => if k < length new_type_names then
+ Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
+ typ_of_dtyp descr sorts dt) $ x
+ else error "nested recursion not (yet) supported"
+ | _ => x) xs :: r_args)
+ end
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
+ val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
+ val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
+ val constrT = map fastype_of l_args ---> T;
+ val lhs = list_comb (Const (cname, constrT), l_args);
+ val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
+ val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (rep_name, T --> T') $ lhs, rhs));
+ val def_name = (Long_Name.base_name cname) ^ "_def";
+ val ([def_thm], thy') = thy |>
+ Sign.add_consts_i [(Binding.name cname', constrT, mx)] |>
+ (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
+ in (thy', defs @ [def_thm], eqns @ [eqn]) end;
+
+ fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
+ (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
+ let
+ val rep_const = cterm_of thy
+ (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
+ val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+ val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
+ ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
+ in
+ (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
+ end;
+
+ val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
+ ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
+ List.take (pdescr, length new_type_names) ~~
+ new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
+
+ val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
+ val rep_inject_thms = map (#Rep_inject o fst) typedefs
+
+ (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *)
+
+ fun prove_constr_rep_thm eqn =
+ let
+ val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
+ val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
+ in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
+ [resolve_tac inj_thms 1,
+ rewrite_goals_tac rewrites,
+ rtac refl 3,
+ resolve_tac rep_intrs 2,
+ REPEAT (resolve_tac Rep_thms 1)])
+ end;
+
+ val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
+
+ (* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
+
+ fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
+ let
+ val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
+ val Type ("fun", [T, U]) = fastype_of Rep;
+ val permT = mk_permT (Type (atom, []));
+ val pi = Free ("pi", permT);
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8
+ (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
+ Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
+ (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
+ perm_closed_thms @ Rep_thms)) 1)
+ end) Rep_thms;
+
+ val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
+ (atoms ~~ perm_closed_thmss));
+
+ (* prove distinctness theorems *)
+
+ val distinct_props = DatatypeProp.make_distincts descr' sorts;
+ val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
+ dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
+ constr_rep_thmss dist_lemmas;
+
+ fun prove_distinct_thms _ (_, []) = []
+ | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) =
+ let
+ val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
+ simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
+ in dist_thm :: standard (dist_thm RS not_sym) ::
+ prove_distinct_thms p (k, ts)
+ end;
+
+ val distinct_thms = map2 prove_distinct_thms
+ (constr_rep_thmss ~~ dist_lemmas) distinct_props;
+
+ (** prove equations for permutation functions **)
+
+ val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
+ let val T = nth_dtyp' i
+ in List.concat (map (fn (atom, perm_closed_thms) =>
+ map (fn ((cname, dts), constr_rep_thm) =>
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+ val permT = mk_permT (Type (atom, []));
+ val pi = Free ("pi", permT);
+
+ fun perm t =
+ let val T = fastype_of t
+ in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
+
+ fun constr_arg ((dts, dt), (j, l_args, r_args)) =
+ let
+ val Ts = map (typ_of_dtyp descr'' sorts) dts;
+ val xs = map (fn (T, i) => mk_Free "x" T i)
+ (Ts ~~ (j upto j + length dts - 1))
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ x :: l_args,
+ map perm (xs @ [x]) @ r_args)
+ end
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts;
+ val c = Const (cname, map fastype_of l_args ---> T)
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8
+ (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
+ (fn _ => EVERY
+ [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
+ simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
+ constr_defs @ perm_closed_thms)) 1,
+ TRY (simp_tac (HOL_basic_ss addsimps
+ (symmetric perm_fun_def :: abs_perm)) 1),
+ TRY (simp_tac (HOL_basic_ss addsimps
+ (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
+ perm_closed_thms)) 1)])
+ end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
+
+ (** prove injectivity of constructors **)
+
+ val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
+ val alpha = PureThy.get_thms thy8 "alpha";
+ val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
+
+ val pt_cp_sort =
+ map (pt_class_of thy8) dt_atoms @
+ maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
+
+ val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
+ let val T = nth_dtyp' i
+ in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
+ if null dts then NONE else SOME
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+
+ fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
+ let
+ val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
+ val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
+ val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ (x :: args1), ys @ (y :: args2),
+ HOLogic.mk_eq
+ (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs)
+ end;
+
+ val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts;
+ val Ts = map fastype_of args1;
+ val c = Const (cname, Ts ---> T)
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
+ foldr1 HOLogic.mk_conj eqs))))
+ (fn _ => EVERY
+ [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
+ rep_inject_thms')) 1,
+ TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
+ alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
+ perm_rep_perm_thms)) 1)])
+ end) (constrs ~~ constr_rep_thms)
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
+
+ (** equations for support and freshness **)
+
+ val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
+ (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
+ let val T = nth_dtyp' i
+ in List.concat (map (fn (cname, dts) => map (fn atom =>
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+ val atomT = Type (atom, []);
+
+ fun process_constr ((dts, dt), (j, args1, args2)) =
+ let
+ val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2)
+ end;
+
+ val (_, args1, args2) = List.foldr process_constr (1, [], []) dts;
+ val Ts = map fastype_of args1;
+ val c = list_comb (Const (cname, Ts ---> T), args1);
+ fun supp t =
+ Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
+ fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
+ val supp_thm = Goal.prove_global thy8 [] []
+ (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (supp c,
+ if null dts then HOLogic.mk_set atomT []
+ else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2)))))
+ (fn _ =>
+ simp_tac (HOL_basic_ss addsimps (supp_def ::
+ Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
+ symmetric empty_def :: finite_emptyI :: simp_thms @
+ abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
+ in
+ (supp_thm,
+ Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fresh c,
+ if null dts then HOLogic.true_const
+ else foldr1 HOLogic.mk_conj (map fresh args2)))))
+ (fn _ =>
+ simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
+ end) atoms) constrs)
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
+
+ (**** weak induction theorem ****)
+
+ fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
+ let
+ val Rep_t = Const (List.nth (rep_names, i), T --> U) $
+ mk_Free "x" T i;
+
+ val Abs_t = Const (List.nth (abs_names, i), U --> T)
+
+ in (prems @ [HOLogic.imp $
+ (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
+ (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
+ concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
+ end;
+
+ val (indrule_lemma_prems, indrule_lemma_concls) =
+ Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
+
+ val indrule_lemma = Goal.prove_global thy8 [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
+ [REPEAT (etac conjE 1),
+ REPEAT (EVERY
+ [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
+ etac mp 1, resolve_tac Rep_thms 1])]);
+
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
+ map (Free o apfst fst o dest_Var) Ps;
+ val indrule_lemma' = cterm_instantiate
+ (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
+
+ val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
+
+ val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
+ val dt_induct = Goal.prove_global thy8 []
+ (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
+ (fn {prems, ...} => EVERY
+ [rtac indrule_lemma' 1,
+ (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ EVERY (map (fn (prem, r) => (EVERY
+ [REPEAT (eresolve_tac Abs_inverse_thms' 1),
+ simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
+ DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
+ (prems ~~ constr_defs))]);
+
+ val case_names_induct = mk_case_names_induct descr'';
+
+ (**** prove that new datatypes have finite support ****)
+
+ val _ = warning "proving finite support for the new datatype";
+
+ val indnames = DatatypeProp.make_tnames recTs;
+
+ val abs_supp = PureThy.get_thms thy8 "abs_supp";
+ val supp_atm = PureThy.get_thms thy8 "supp_atm";
+
+ val finite_supp_thms = map (fn atom =>
+ let val atomT = Type (atom, [])
+ in map standard (List.take
+ (split_conj_thm (Goal.prove_global thy8 [] []
+ (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
+ (HOLogic.mk_Trueprop
+ (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
+ Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
+ (indnames ~~ recTs)))))
+ (fn _ => indtac dt_induct indnames 1 THEN
+ ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
+ (abs_supp @ supp_atm @
+ PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
+ List.concat supp_thms))))),
+ length new_type_names))
+ end) atoms;
+
+ val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
+
+ (* Function to add both the simp and eqvt attributes *)
+ (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
+
+ val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
+
+ val (_, thy9) = thy8 |>
+ Sign.add_path big_name |>
+ PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
+ PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
+ Sign.parent_path ||>>
+ DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
+ DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
+ DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
+ DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
+ DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
+ DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
+ fold (fn (atom, ths) => fn thy =>
+ let
+ val class = fs_class_of thy atom;
+ val sort = Sign.certify_sort thy (class :: pt_cp_sort)
+ in fold (fn Type (s, Ts) => AxClass.prove_arity
+ (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
+ (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
+ end) (atoms ~~ finite_supp_thms);
+
+ (**** strong induction theorem ****)
+
+ val pnames = if length descr'' = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
+ val ind_sort = if null dt_atomTs then HOLogic.typeS
+ else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
+ val fsT = TFree ("'n", ind_sort);
+ val fsT' = TFree ("'n", HOLogic.typeS);
+
+ val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
+ (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
+
+ fun make_pred fsT i T =
+ Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
+
+ fun mk_fresh1 xs [] = []
+ | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
+ (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
+ (filter (fn (_, U) => T = U) (rev xs)) @
+ mk_fresh1 (y :: xs) ys;
+
+ fun mk_fresh2 xss [] = []
+ | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
+ map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
+ (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @
+ mk_fresh2 (p :: xss) yss;
+
+ fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
+ let
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr'' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr'' sorts) recs;
+ val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = tnames ~~ Ts;
+ val frees' = partition_cargs idxs frees;
+ val z = (Name.variant tnames "z", fsT);
+
+ fun mk_prem ((dt, s), T) =
+ let
+ val (Us, U) = strip_type T;
+ val l = length Us
+ in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
+ end;
+
+ val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
+ val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
+ (f T (Free p) (Free z))) (List.concat (map fst frees')) @
+ mk_fresh1 [] (List.concat (map fst frees')) @
+ mk_fresh2 [] frees'
+
+ in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
+ HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
+ list_comb (Const (cname, Ts ---> T), map Free frees))))
+ end;
+
+ val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
+ map (make_ind_prem fsT (fn T => fn t => fn u =>
+ fresh_const T fsT $ t $ u) i T)
+ (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
+ val tnames = DatatypeProp.make_tnames recTs;
+ val zs = Name.variant_list tnames (replicate (length descr'') "z");
+ val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), tname), z) =>
+ make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
+ (descr'' ~~ recTs ~~ tnames ~~ zs)));
+ val induct = Logic.list_implies (ind_prems, ind_concl);
+
+ val ind_prems' =
+ map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
+ HOLogic.mk_Trueprop (Const ("Finite_Set.finite",
+ (snd (split_last (binder_types T)) --> HOLogic.boolT) -->
+ HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
+ List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
+ map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
+ HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
+ (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
+ val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), tname), z) =>
+ make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
+ (descr'' ~~ recTs ~~ tnames ~~ zs)));
+ val induct' = Logic.list_implies (ind_prems', ind_concl');
+
+ val aux_ind_vars =
+ (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
+ map mk_permT dt_atomTs) @ [("z", fsT')];
+ val aux_ind_Ts = rev (map snd aux_ind_vars);
+ val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn (((i, _), T), tname) =>
+ HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
+ fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
+ (Free (tname, T))))
+ (descr'' ~~ recTs ~~ tnames)));
+
+ val fin_set_supp = map (fn s =>
+ at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
+ val fin_set_fresh = map (fn s =>
+ at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
+ val pt1_atoms = map (fn Type (s, _) =>
+ PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
+ val pt2_atoms = map (fn Type (s, _) =>
+ PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
+ val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
+ val fs_atoms = PureThy.get_thms thy9 "fin_supp";
+ val abs_supp = PureThy.get_thms thy9 "abs_supp";
+ val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh";
+ val calc_atm = PureThy.get_thms thy9 "calc_atm";
+ val fresh_atm = PureThy.get_thms thy9 "fresh_atm";
+ val fresh_left = PureThy.get_thms thy9 "fresh_left";
+ val perm_swap = PureThy.get_thms thy9 "perm_swap";
+
+ fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
+ let
+ val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
+ val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
+ (HOLogic.exists_const T $ Abs ("x", T,
+ fresh_const T (fastype_of p) $
+ Bound 0 $ p)))
+ (fn _ => EVERY
+ [resolve_tac exists_fresh' 1,
+ simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
+ fin_set_supp @ ths)) 1]);
+ val (([cx], ths), ctxt') = Obtain.result
+ (fn _ => EVERY
+ [etac exE 1,
+ full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
+ REPEAT (etac conjE 1)])
+ [ex] ctxt
+ in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+
+ fun fresh_fresh_inst thy a b =
+ let
+ val T = fastype_of a;
+ val SOME th = find_first (fn th => case prop_of th of
+ _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
+ | _ => false) perm_fresh_fresh
+ in
+ Drule.instantiate' []
+ [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
+ end;
+
+ val fs_cp_sort =
+ map (fs_class_of thy9) dt_atoms @
+ maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
+
+ (**********************************************************************
+ The subgoals occurring in the proof of induct_aux have the
+ following parameters:
+
+ x_1 ... x_k p_1 ... p_m z
+
+ where
+
+ x_i : constructor arguments (introduced by weak induction rule)
+ p_i : permutations (one for each atom type in the data type)
+ z : freshness context
+ ***********************************************************************)
+
+ val _ = warning "proving strong induction theorem ...";
+
+ val induct_aux = Goal.prove_global thy9 []
+ (map (augment_sort thy9 fs_cp_sort) ind_prems')
+ (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
+ let
+ val (prems1, prems2) = chop (length dt_atomTs) prems;
+ val ind_ss2 = HOL_ss addsimps
+ finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
+ val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
+ fresh_atm @ rev_simps @ app_simps;
+ val ind_ss3 = HOL_ss addsimps abs_fun_eq1 ::
+ abs_perm @ calc_atm @ perm_swap;
+ val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
+ fin_set_fresh @ calc_atm;
+ val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
+ val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
+ val th = Goal.prove context [] []
+ (augment_sort thy9 fs_cp_sort aux_ind_concl)
+ (fn {context = context1, ...} =>
+ EVERY (indtac dt_induct tnames 1 ::
+ maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
+ map (fn ((cname, cargs), is) =>
+ REPEAT (rtac allI 1) THEN
+ SUBPROOF (fn {prems = iprems, params, concl,
+ context = context2, ...} =>
+ let
+ val concl' = term_of concl;
+ val _ $ (_ $ _ $ u) = concl';
+ val U = fastype_of u;
+ val (xs, params') =
+ chop (length cargs) (map term_of params);
+ val Ts = map fastype_of xs;
+ val cnstr = Const (cname, Ts ---> U);
+ val (pis, z) = split_last params';
+ val mk_pi = fold_rev (mk_perm []) pis;
+ val xs' = partition_cargs is xs;
+ val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
+ val ts = maps (fn (ts, u) => ts @ [u]) xs'';
+ val (freshs1, freshs2, context3) = fold (fn t =>
+ let val T = fastype_of t
+ in obtain_fresh_name' prems1
+ (the (AList.lookup op = fresh_fs T) $ z :: ts) T
+ end) (maps fst xs') ([], [], context2);
+ val freshs1' = unflat (map fst xs') freshs1;
+ val freshs2' = map (Simplifier.simplify ind_ss4)
+ (mk_not_sym freshs2);
+ val ind_ss1' = ind_ss1 addsimps freshs2';
+ val ind_ss3' = ind_ss3 addsimps freshs2';
+ val rename_eq =
+ if forall (null o fst) xs' then []
+ else [Goal.prove context3 [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (list_comb (cnstr, ts),
+ list_comb (cnstr, maps (fn ((bs, t), cs) =>
+ cs @ [fold_rev (mk_perm []) (map perm_of_pair
+ (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
+ (fn _ => EVERY
+ (simp_tac (HOL_ss addsimps flat inject_thms) 1 ::
+ REPEAT (FIRSTGOAL (rtac conjI)) ::
+ maps (fn ((bs, t), cs) =>
+ if null bs then []
+ else rtac sym 1 :: maps (fn (b, c) =>
+ [rtac trans 1, rtac sym 1,
+ rtac (fresh_fresh_inst thy9 b c) 1,
+ simp_tac ind_ss1' 1,
+ simp_tac ind_ss2 1,
+ simp_tac ind_ss3' 1]) (bs ~~ cs))
+ (xs'' ~~ freshs1')))];
+ val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
+ [simp_tac (ind_ss6 addsimps rename_eq) 1,
+ cut_facts_tac iprems 1,
+ (resolve_tac prems THEN_ALL_NEW
+ SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
+ _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
+ simp_tac ind_ss1' i
+ | _ $ (Const ("Not", _) $ _) =>
+ resolve_tac freshs2' i
+ | _ => asm_simp_tac (HOL_basic_ss addsimps
+ pt2_atoms addsimprocs [perm_simproc]) i)) 1])
+ val final = ProofContext.export context3 context2 [th]
+ in
+ resolve_tac final 1
+ end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
+ in
+ EVERY
+ [cut_facts_tac [th] 1,
+ REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
+ REPEAT (etac allE 1),
+ REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
+ end);
+
+ val induct_aux' = Thm.instantiate ([],
+ map (fn (s, v as Var (_, T)) =>
+ (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
+ (pnames ~~ map head_of (HOLogic.dest_conj
+ (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
+ map (fn (_, f) =>
+ let val f' = Logic.varify f
+ in (cterm_of thy9 f',
+ cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
+ end) fresh_fs) induct_aux;
+
+ val induct = Goal.prove_global thy9 []
+ (map (augment_sort thy9 fs_cp_sort) ind_prems)
+ (augment_sort thy9 fs_cp_sort ind_concl)
+ (fn {prems, ...} => EVERY
+ [rtac induct_aux' 1,
+ REPEAT (resolve_tac fs_atoms 1),
+ REPEAT ((resolve_tac prems THEN_ALL_NEW
+ (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
+
+ val (_, thy10) = thy9 |>
+ Sign.add_path big_name |>
+ PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
+ PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
+ PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
+
+ (**** recursion combinator ****)
+
+ val _ = warning "defining recursion combinator ...";
+
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+
+ val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
+
+ val rec_sort = if null dt_atomTs then HOLogic.typeS else
+ Sign.certify_sort thy10 pt_cp_sort;
+
+ val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
+ val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
+
+ val rec_set_Ts = map (fn (T1, T2) =>
+ rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
+
+ val big_rec_name = big_name ^ "_rec_set";
+ val rec_set_names' =
+ if length descr'' = 1 then [big_rec_name] else
+ map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
+ (1 upto (length descr''));
+ val rec_set_names = map (Sign.full_bname thy10) rec_set_names';
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
+ val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
+ (rec_set_names' ~~ rec_set_Ts);
+ val rec_sets = map (fn c => list_comb (Const c, rec_fns))
+ (rec_set_names ~~ rec_set_Ts);
+
+ (* introduction rules for graph of recursion function *)
+
+ val rec_preds = map (fn (a, T) =>
+ Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
+
+ fun mk_fresh3 rs [] = []
+ | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
+ List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
+ else SOME (HOLogic.mk_Trueprop
+ (fresh_const T U $ Free y $ Free r))) rs) ys) @
+ mk_fresh3 rs yss;
+
+ (* FIXME: avoid collisions with other variable names? *)
+ val rec_ctxt = Free ("z", fsT');
+
+ fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
+ rec_eq_prems, l), ((cname, cargs), idxs)) =
+ let
+ val Ts = map (typ_of_dtyp descr'' sorts) cargs;
+ val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
+ val frees' = partition_cargs idxs frees;
+ val binders = List.concat (map fst frees');
+ val atomTs = distinct op = (maps (map snd o fst) frees');
+ val recs = List.mapPartial
+ (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
+ (partition_cargs idxs cargs ~~ frees');
+ val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
+ map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
+ val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
+ (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
+ val prems2 =
+ map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
+ (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
+ val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
+ val prems4 = map (fn ((i, _), y) =>
+ HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
+ val prems5 = mk_fresh3 (recs ~~ frees'') frees';
+ val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
+ frees'') atomTs;
+ val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
+ (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
+ val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
+ val result_freshs = map (fn p as (_, T) =>
+ fresh_const T (fastype_of result) $ Free p $ result) binders;
+ val P = HOLogic.mk_Trueprop (p $ result)
+ in
+ (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
+ HOLogic.mk_Trueprop (rec_set $
+ list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
+ rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
+ rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
+ Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
+ HOLogic.mk_Trueprop fr))) result_freshs,
+ rec_eq_prems @ [List.concat prems2 @ prems3],
+ l + 1)
+ end;
+
+ val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
+ Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
+ Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
+ (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
+
+ val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
+ thy10 |>
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+ (map dest_Free rec_fns)
+ (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
+ PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct");
+
+ (** equivariance **)
+
+ val fresh_bij = PureThy.get_thms thy11 "fresh_bij";
+ val perm_bij = PureThy.get_thms thy11 "perm_bij";
+
+ val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
+ let
+ val permT = mk_permT aT;
+ val pi = Free ("pi", permT);
+ val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
+ (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
+ val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
+ (rec_set_names ~~ rec_set_Ts);
+ val ps = map (fn ((((T, U), R), R'), i) =>
+ let
+ val x = Free ("x" ^ string_of_int i, T);
+ val y = Free ("y" ^ string_of_int i, U)
+ in
+ (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
+ val ths = map (fn th => standard (th RS mp)) (split_conj_thm
+ (Goal.prove_global thy11 [] []
+ (augment_sort thy1 pt_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
+ (fn _ => rtac rec_induct 1 THEN REPEAT
+ (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
+ addsimps flat perm_simps'
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
+ (resolve_tac rec_intrs THEN_ALL_NEW
+ asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
+ val ths' = map (fn ((P, Q), th) =>
+ Goal.prove_global thy11 [] []
+ (augment_sort thy1 pt_cp_sort
+ (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
+ (fn _ => dtac (Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), permT)),
+ cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
+ NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
+ in (ths, ths') end) dt_atomTs);
+
+ (** finite support **)
+
+ val rec_fin_supp_thms = map (fn aT =>
+ let
+ val name = Long_Name.base_name (fst (dest_Type aT));
+ val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
+ val aset = HOLogic.mk_setT aT;
+ val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
+ val fins = map (fn (f, T) => HOLogic.mk_Trueprop
+ (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
+ (rec_fns ~~ rec_fn_Ts)
+ in
+ map (fn th => standard (th RS mp)) (split_conj_thm
+ (Goal.prove_global thy11 []
+ (map (augment_sort thy11 fs_cp_sort) fins)
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((T, U), R), i) =>
+ let
+ val x = Free ("x" ^ string_of_int i, T);
+ val y = Free ("y" ^ string_of_int i, U)
+ in
+ HOLogic.mk_imp (R $ x $ y,
+ finite $ (Const ("Nominal.supp", U --> aset) $ y))
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
+ (1 upto length recTs))))))
+ (fn {prems = fins, ...} =>
+ (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
+ (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
+ end) dt_atomTs;
+
+ (** freshness **)
+
+ val finite_premss = map (fn aT =>
+ map (fn (f, T) => HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
+ (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
+
+ val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
+
+ val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
+ let
+ val name = Long_Name.base_name (fst (dest_Type aT));
+ val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
+ val a = Free ("a", aT);
+ val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
+ (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
+ in
+ map (fn (((T, U), R), eqvt_th) =>
+ let
+ val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
+ val y = Free ("y", U);
+ val y' = Free ("y'", U)
+ in
+ standard (Goal.prove (ProofContext.init thy11) []
+ (map (augment_sort thy11 fs_cp_sort)
+ (finite_prems @
+ [HOLogic.mk_Trueprop (R $ x $ y),
+ HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
+ HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
+ HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
+ freshs))
+ (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
+ (fn {prems, context} =>
+ let
+ val (finite_prems, rec_prem :: unique_prem ::
+ fresh_prems) = chop (length finite_prems) prems;
+ val unique_prem' = unique_prem RS spec RS mp;
+ val unique = [unique_prem', unique_prem' RS sym] MRS trans;
+ val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
+ val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
+ in EVERY
+ [rtac (Drule.cterm_instantiate
+ [(cterm_of thy11 S,
+ cterm_of thy11 (Const ("Nominal.supp",
+ fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
+ supports_fresh) 1,
+ simp_tac (HOL_basic_ss addsimps
+ [supports_def, symmetric fresh_def, fresh_prod]) 1,
+ REPEAT_DETERM (resolve_tac [allI, impI] 1),
+ REPEAT_DETERM (etac conjE 1),
+ rtac unique 1,
+ SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
+ [cut_facts_tac [rec_prem] 1,
+ rtac (Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
+ cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
+ asm_simp_tac (HOL_ss addsimps
+ (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
+ rtac rec_prem 1,
+ simp_tac (HOL_ss addsimps (fs_name ::
+ supp_prod :: finite_Un :: finite_prems)) 1,
+ simp_tac (HOL_ss addsimps (symmetric fresh_def ::
+ fresh_prod :: fresh_prems)) 1]
+ end))
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
+ end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
+
+ (** uniqueness **)
+
+ val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
+ val fun_tupleT = fastype_of fun_tuple;
+ val rec_unique_frees =
+ DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
+ val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
+ val rec_unique_frees' =
+ DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
+ val rec_unique_concls = map (fn ((x, U), R) =>
+ Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
+ Abs ("y", U, R $ Free x $ Bound 0))
+ (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
+
+ val induct_aux_rec = Drule.cterm_instantiate
+ (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
+ (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
+ Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
+ fresh_fs @
+ map (fn (((P, T), (x, U)), Q) =>
+ (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
+ Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
+ (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
+ map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
+ rec_unique_frees)) induct_aux;
+
+ fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
+ let
+ val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
+ val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
+ (HOLogic.exists_const T $ Abs ("x", T,
+ fresh_const T (fastype_of p) $ Bound 0 $ p)))
+ (fn _ => EVERY
+ [cut_facts_tac ths 1,
+ REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
+ resolve_tac exists_fresh' 1,
+ asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
+ val (([cx], ths), ctxt') = Obtain.result
+ (fn _ => EVERY
+ [etac exE 1,
+ full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
+ REPEAT (etac conjE 1)])
+ [ex] ctxt
+ in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+
+ val finite_ctxt_prems = map (fn aT =>
+ HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
+
+ val rec_unique_thms = split_conj_thm (Goal.prove
+ (ProofContext.init thy11) (map fst rec_unique_frees)
+ (map (augment_sort thy11 fs_cp_sort)
+ (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
+ (fn {prems, context} =>
+ let
+ val k = length rec_fns;
+ val (finite_thss, ths1) = fold_map (fn T => fn xs =>
+ apfst (pair T) (chop k xs)) dt_atomTs prems;
+ val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
+ val (P_ind_ths, fcbs) = chop k ths2;
+ val P_ths = map (fn th => th RS mp) (split_conj_thm
+ (Goal.prove context
+ (map fst (rec_unique_frees'' @ rec_unique_frees')) []
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((x, y), S), P) => HOLogic.mk_imp
+ (S $ Free x $ Free y, P $ (Free y)))
+ (rec_unique_frees'' ~~ rec_unique_frees' ~~
+ rec_sets ~~ rec_preds)))))
+ (fn _ =>
+ rtac rec_induct 1 THEN
+ REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
+ val rec_fin_supp_thms' = map
+ (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
+ (rec_fin_supp_thms ~~ finite_thss);
+ in EVERY
+ ([rtac induct_aux_rec 1] @
+ maps (fn ((_, finite_ths), finite_th) =>
+ [cut_facts_tac (finite_th :: finite_ths) 1,
+ asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
+ (finite_thss ~~ finite_ctxt_ths) @
+ maps (fn ((_, idxss), elim) => maps (fn idxs =>
+ [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
+ REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
+ rtac ex1I 1,
+ (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
+ rotate_tac ~1 1,
+ ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
+ (HOL_ss addsimps List.concat distinct_thms)) 1] @
+ (if null idxs then [] else [hyp_subst_tac 1,
+ SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
+ let
+ val SOME prem = find_first (can (HOLogic.dest_eq o
+ HOLogic.dest_Trueprop o prop_of)) prems';
+ val _ $ (_ $ lhs $ rhs) = prop_of prem;
+ val _ $ (_ $ lhs' $ rhs') = term_of concl;
+ val rT = fastype_of lhs';
+ val (c, cargsl) = strip_comb lhs;
+ val cargsl' = partition_cargs idxs cargsl;
+ val boundsl = List.concat (map fst cargsl');
+ val (_, cargsr) = strip_comb rhs;
+ val cargsr' = partition_cargs idxs cargsr;
+ val boundsr = List.concat (map fst cargsr');
+ val (params1, _ :: params2) =
+ chop (length params div 2) (map term_of params);
+ val params' = params1 @ params2;
+ val rec_prems = filter (fn th => case prop_of th of
+ _ $ p => (case head_of p of
+ Const (s, _) => s mem rec_set_names
+ | _ => false)
+ | _ => false) prems';
+ val fresh_prems = filter (fn th => case prop_of th of
+ _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
+ | _ $ (Const ("Not", _) $ _) => true
+ | _ => false) prems';
+ val Ts = map fastype_of boundsl;
+
+ val _ = warning "step 1: obtaining fresh names";
+ val (freshs1, freshs2, context'') = fold
+ (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
+ (List.concat (map snd finite_thss) @
+ finite_ctxt_ths @ rec_prems)
+ rec_fin_supp_thms')
+ Ts ([], [], context');
+ val pi1 = map perm_of_pair (boundsl ~~ freshs1);
+ val rpi1 = rev pi1;
+ val pi2 = map perm_of_pair (boundsr ~~ freshs1);
+ val rpi2 = rev pi2;
+
+ val fresh_prems' = mk_not_sym fresh_prems;
+ val freshs2' = mk_not_sym freshs2;
+
+ (** as, bs, cs # K as ts, K bs us **)
+ val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
+ val prove_fresh_ss = HOL_ss addsimps
+ (finite_Diff :: List.concat fresh_thms @
+ fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
+ (* FIXME: avoid asm_full_simp_tac ? *)
+ fun prove_fresh ths y x = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const
+ (fastype_of x) (fastype_of y) $ x $ y))
+ (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
+ val constr_fresh_thms =
+ map (prove_fresh fresh_prems lhs) boundsl @
+ map (prove_fresh fresh_prems rhs) boundsr @
+ map (prove_fresh freshs2 lhs) freshs1 @
+ map (prove_fresh freshs2 rhs) freshs1;
+
+ (** pi1 o (K as ts) = pi2 o (K bs us) **)
+ val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
+ val pi1_pi2_eq = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
+ (fn _ => EVERY
+ [cut_facts_tac constr_fresh_thms 1,
+ asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
+ rtac prem 1]);
+
+ (** pi1 o ts = pi2 o us **)
+ val _ = warning "step 4: pi1 o ts = pi2 o us";
+ val pi1_pi2_eqs = map (fn (t, u) =>
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
+ (fn _ => EVERY
+ [cut_facts_tac [pi1_pi2_eq] 1,
+ asm_full_simp_tac (HOL_ss addsimps
+ (calc_atm @ List.concat perm_simps' @
+ fresh_prems' @ freshs2' @ abs_perm @
+ alpha @ List.concat inject_thms)) 1]))
+ (map snd cargsl' ~~ map snd cargsr');
+
+ (** pi1^-1 o pi2 o us = ts **)
+ val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
+ val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
+ (fn _ => simp_tac (HOL_ss addsimps
+ ((eq RS sym) :: perm_swap)) 1))
+ (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
+
+ val (rec_prems1, rec_prems2) =
+ chop (length rec_prems div 2) rec_prems;
+
+ (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
+ val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
+ val rec_prems' = map (fn th =>
+ let
+ val _ $ (S $ x $ y) = prop_of th;
+ val Const (s, _) = head_of S;
+ val k = find_index (equal s) rec_set_names;
+ val pi = rpi1 @ pi2;
+ fun mk_pi z = fold_rev (mk_perm []) pi z;
+ fun eqvt_tac p =
+ let
+ val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
+ val l = find_index (equal T) dt_atomTs;
+ val th = List.nth (List.nth (rec_equiv_thms', l), k);
+ val th' = Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), U)),
+ cterm_of thy11 p)]) th;
+ in rtac th' 1 end;
+ val th' = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
+ (fn _ => EVERY
+ (map eqvt_tac pi @
+ [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
+ perm_swap @ perm_fresh_fresh)) 1,
+ rtac th 1]))
+ in
+ Simplifier.simplify
+ (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
+ end) rec_prems2;
+
+ val ihs = filter (fn th => case prop_of th of
+ _ $ (Const ("All", _) $ _) => true | _ => false) prems';
+
+ (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
+ val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
+ val rec_eqns = map (fn (th, ih) =>
+ let
+ val th' = th RS (ih RS spec RS mp) RS sym;
+ val _ $ (_ $ lhs $ rhs) = prop_of th';
+ fun strip_perm (_ $ _ $ t) = strip_perm t
+ | strip_perm t = t;
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 lhs,
+ fold_rev (mk_perm []) pi2 (strip_perm rhs))))
+ (fn _ => simp_tac (HOL_basic_ss addsimps
+ (th' :: perm_swap)) 1)
+ end) (rec_prems' ~~ ihs);
+
+ (** as # rs **)
+ val _ = warning "step 8: as # rs";
+ val rec_freshs = List.concat
+ (map (fn (rec_prem, ih) =>
+ let
+ val _ $ (S $ x $ (y as Free (_, T))) =
+ prop_of rec_prem;
+ val k = find_index (equal S) rec_sets;
+ val atoms = List.concat (List.mapPartial (fn (bs, z) =>
+ if z = x then NONE else SOME bs) cargsl')
+ in
+ map (fn a as Free (_, aT) =>
+ let val l = find_index (equal aT) dt_atomTs;
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
+ (fn _ => EVERY
+ (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
+ map (fn th => rtac th 1)
+ (snd (List.nth (finite_thss, l))) @
+ [rtac rec_prem 1, rtac ih 1,
+ REPEAT_DETERM (resolve_tac fresh_prems 1)]))
+ end) atoms
+ end) (rec_prems1 ~~ ihs));
+
+ (** as # fK as ts rs , bs # fK bs us vs **)
+ val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
+ fun prove_fresh_result (a as Free (_, aT)) =
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
+ (fn _ => EVERY
+ [resolve_tac fcbs 1,
+ REPEAT_DETERM (resolve_tac
+ (fresh_prems @ rec_freshs) 1),
+ REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
+ THEN resolve_tac rec_prems 1),
+ resolve_tac P_ind_ths 1,
+ REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
+
+ val fresh_results'' = map prove_fresh_result boundsl;
+
+ fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
+ let val th' = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $
+ fold_rev (mk_perm []) (rpi2 @ pi1) a $
+ fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
+ (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
+ rtac th 1)
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
+ (fn _ => EVERY
+ [cut_facts_tac [th'] 1,
+ full_simp_tac (Simplifier.theory_context thy11 HOL_ss
+ addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1,
+ full_simp_tac (HOL_ss addsimps (calc_atm @
+ fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
+ end;
+
+ val fresh_results = fresh_results'' @ map prove_fresh_result''
+ (boundsl ~~ boundsr ~~ fresh_results'');
+
+ (** cs # fK as ts rs , cs # fK bs us vs **)
+ val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
+ fun prove_fresh_result' recs t (a as Free (_, aT)) =
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
+ (fn _ => EVERY
+ [cut_facts_tac recs 1,
+ REPEAT_DETERM (dresolve_tac
+ (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
+ NominalPermeq.fresh_guess_tac
+ (HOL_ss addsimps (freshs2 @
+ fs_atoms @ fresh_atm @
+ List.concat (map snd finite_thss))) 1]);
+
+ val fresh_results' =
+ map (prove_fresh_result' rec_prems1 rhs') freshs1 @
+ map (prove_fresh_result' rec_prems2 lhs') freshs1;
+
+ (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
+ val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
+ val pi1_pi2_result = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
+ (fn _ => simp_tac (Simplifier.context context'' HOL_ss
+ addsimps pi1_pi2_eqs @ rec_eqns
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
+ TRY (simp_tac (HOL_ss addsimps
+ (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
+
+ val _ = warning "final result";
+ val final = Goal.prove context'' [] [] (term_of concl)
+ (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
+ full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
+ fresh_results @ fresh_results') 1);
+ val final' = ProofContext.export context'' context' [final];
+ val _ = warning "finished!"
+ in
+ resolve_tac final' 1
+ end) context 1])) idxss) (ndescr ~~ rec_elims))
+ end));
+
+ val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
+
+ (* define primrec combinators *)
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.full_bname thy11)
+ (if length descr'' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr''))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ val (reccomb_defs, thy12) =
+ thy11
+ |> Sign.add_consts_i (map (fn ((name, T), T') =>
+ (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts))
+ |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
+ (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
+ Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+ set $ Free ("x", T) $ Free ("y", T'))))))
+ (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
+
+ (* prove characteristic equations for primrec combinators *)
+
+ val rec_thms = map (fn (prems, concl) =>
+ let
+ val _ $ (_ $ (_ $ x) $ _) = concl;
+ val (_, cargs) = strip_comb x;
+ val ps = map (fn (x as Free (_, T), i) =>
+ (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
+ val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
+ val prems' = List.concat finite_premss @ finite_ctxt_prems @
+ rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
+ fun solve rules prems = resolve_tac rules THEN_ALL_NEW
+ (resolve_tac prems THEN_ALL_NEW atac)
+ in
+ Goal.prove_global thy12 []
+ (map (augment_sort thy12 fs_cp_sort) prems')
+ (augment_sort thy12 fs_cp_sort concl')
+ (fn {prems, ...} => EVERY
+ [rewrite_goals_tac reccomb_defs,
+ rtac the1_equality 1,
+ solve rec_unique_thms prems 1,
+ resolve_tac rec_intrs 1,
+ REPEAT (solve (prems @ rec_total_thms) prems 1)])
+ end) (rec_eq_prems ~~
+ DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
+
+ val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
+ ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
+
+ (* FIXME: theorems are stored in database for testing only *)
+ val (_, thy13) = thy12 |>
+ PureThy.add_thmss
+ [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
+ ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
+ ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
+ ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
+ ((Binding.name "rec_unique", map standard rec_unique_thms), []),
+ ((Binding.name "recs", rec_thms), [])] ||>
+ Sign.parent_path ||>
+ map_nominal_datatypes (fold Symtab.update dt_infos);
+
+ in
+ thy13
+ end;
+
+val add_nominal_datatype = gen_add_nominal_datatype Datatype.read_typ;
+
+
+(* FIXME: The following stuff should be exported by Datatype *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val datatype_decl =
+ Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
+ (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
+
+fun mk_datatype args =
+ let
+ val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
+ val specs = map (fn ((((_, vs), t), mx), cons) =>
+ (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
+ in add_nominal_datatype Datatype.default_config names specs end;
+
+val _ =
+ OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
+ (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
+
+end;
+
+end
--- a/src/HOL/Nominal/nominal_atoms.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML Tue Jun 23 15:48:55 2009 +0100
@@ -101,8 +101,8 @@
val (_,thy1) =
fold_map (fn ak => fn thy =>
let val dt = ([], Binding.name ak, NoSyn, [(Binding.name ak, [@{typ nat}], NoSyn)])
- val ({inject,case_thms,...},thy1) = DatatypePackage.add_datatype
- DatatypeAux.default_datatype_config [ak] [dt] thy
+ val ({inject,case_thms,...},thy1) = Datatype.add_datatype
+ Datatype.default_config [ak] [dt] thy
val inject_flat = flat inject
val ak_type = Type (Sign.intern_type thy1 ak,[])
val ak_sign = Sign.intern_const thy1 ak
@@ -191,7 +191,7 @@
thy |> Sign.add_consts_i [(Binding.name ("swap_" ^ ak_name), swapT, NoSyn)]
|> PureThy.add_defs_unchecked true [((Binding.name name, def2),[])]
|> snd
- |> OldPrimrecPackage.add_primrec_unchecked_i "" [(("", def1),[])]
+ |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1),[])]
end) ak_names_types thy1;
(* declares a permutation function for every atom-kind acting *)
@@ -219,7 +219,7 @@
Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a)));
in
thy |> Sign.add_consts_i [(Binding.name prm_name, mk_permT T --> T --> T, NoSyn)]
- |> OldPrimrecPackage.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])]
+ |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])]
end) ak_names_types thy3;
(* defines permutation functions for all combinations of atom-kinds; *)
--- a/src/HOL/Nominal/nominal_inductive.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML Tue Jun 23 15:48:55 2009 +0100
@@ -53,7 +53,7 @@
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
(Const (s, T), ts) => (case strip_type T of
(Ts, Type (tname, _)) =>
- (case NominalPackage.get_nominal_datatype thy tname of
+ (case Nominal.get_nominal_datatype thy tname of
NONE => fold (add_binders thy i) ts bs
| SOME {descr, index, ...} => (case AList.lookup op =
(#3 (the (AList.lookup op = descr index))) s of
@@ -148,11 +148,11 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
- val ind_params = InductivePackage.params_of raw_induct;
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
+ val ind_params = Inductive.params_of raw_induct;
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
- val monos = InductivePackage.get_monos ctxt;
+ val monos = Inductive.get_monos ctxt;
val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
[] => ()
@@ -230,7 +230,7 @@
else NONE) xs @ mk_distinct xs;
fun mk_fresh (x, T) = HOLogic.mk_Trueprop
- (NominalPackage.fresh_const T fsT $ x $ Bound 0);
+ (Nominal.fresh_const T fsT $ x $ Bound 0);
val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) =>
let
@@ -254,7 +254,7 @@
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
HOLogic.list_all (ind_vars, lift_pred p
- (map (fold_rev (NominalPackage.mk_perm ind_Ts)
+ (map (fold_rev (Nominal.mk_perm ind_Ts)
(map Bound (length atomTs downto 1))) ts)))) concls));
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -268,7 +268,7 @@
else map_term (split_conj (K o I) names) prem prem) prems, q))))
(mk_distinct bvars @
maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
- (NominalPackage.fresh_const U T $ u $ t)) bvars)
+ (Nominal.fresh_const U T $ u $ t)) bvars)
(ts ~~ binder_types (fastype_of p)))) prems;
val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp";
@@ -296,7 +296,7 @@
val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1);
val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
(HOLogic.exists_const T $ Abs ("x", T,
- NominalPackage.fresh_const T (fastype_of p) $
+ Nominal.fresh_const T (fastype_of p) $
Bound 0 $ p)))
(fn _ => EVERY
[resolve_tac exists_fresh' 1,
@@ -325,13 +325,13 @@
(fn (Bound i, T) => (nth params' (length params' - i), T)
| (t, T) => (t, T)) bvars;
val pi_bvars = map (fn (t, _) =>
- fold_rev (NominalPackage.mk_perm []) pis t) bvars';
+ fold_rev (Nominal.mk_perm []) pis t) bvars';
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
val (freshs1, freshs2, ctxt'') = fold
(obtain_fresh_name (ts @ pi_bvars))
(map snd bvars') ([], [], ctxt');
- val freshs2' = NominalPackage.mk_not_sym freshs2;
- val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1);
+ val freshs2' = Nominal.mk_not_sym freshs2;
+ val pis' = map Nominal.perm_of_pair (pi_bvars ~~ freshs1);
fun concat_perm pi1 pi2 =
let val T = fastype_of pi1
in if T = fastype_of pi2 then
@@ -343,11 +343,11 @@
(Vartab.empty, Vartab.empty);
val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy))
(map (Envir.subst_vars env) vs ~~
- map (fold_rev (NominalPackage.mk_perm [])
+ map (fold_rev (Nominal.mk_perm [])
(rev pis' @ pis)) params' @ [z])) ihyp;
fun mk_pi th =
Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
- addsimprocs [NominalPackage.perm_simproc])
+ addsimprocs [Nominal.perm_simproc])
(Simplifier.simplify eqvt_ss
(fold_rev (mk_perm_bool o cterm_of thy)
(rev pis' @ pis) th));
@@ -369,13 +369,13 @@
| _ $ (_ $ (_ $ lhs $ rhs)) =>
(curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs));
val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop
- (bop (fold_rev (NominalPackage.mk_perm []) pis lhs)
- (fold_rev (NominalPackage.mk_perm []) pis rhs)))
+ (bop (fold_rev (Nominal.mk_perm []) pis lhs)
+ (fold_rev (Nominal.mk_perm []) pis rhs)))
(fn _ => simp_tac (HOL_basic_ss addsimps
(fresh_bij @ perm_bij)) 1 THEN rtac th' 1)
in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end)
vc_compat_ths;
- val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths';
+ val vc_compat_ths'' = Nominal.mk_not_sym vc_compat_ths';
(** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **)
(** we have to pre-simplify the rewrite rules **)
val swap_simps_ss = HOL_ss addsimps swap_simps @
@@ -383,14 +383,14 @@
(vc_compat_ths'' @ freshs2');
val th = Goal.prove ctxt'' [] []
(HOLogic.mk_Trueprop (list_comb (P $ hd ts,
- map (fold (NominalPackage.mk_perm []) pis') (tl ts))))
+ map (fold (Nominal.mk_perm []) pis') (tl ts))))
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1,
REPEAT_DETERM_N (nprems_of ihyp - length gprems)
(simp_tac swap_simps_ss 1),
REPEAT_DETERM_N (length gprems)
(simp_tac (HOL_basic_ss
addsimps [inductive_forall_def']
- addsimprocs [NominalPackage.perm_simproc]) 1 THEN
+ addsimprocs [Nominal.perm_simproc]) 1 THEN
resolve_tac gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
@@ -435,7 +435,7 @@
((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt')
end) (prems ~~ avoids) ctxt')
end)
- (InductivePackage.partition_rules' raw_induct (intrs ~~ avoids') ~~
+ (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~
elims);
val cases_prems' =
@@ -448,7 +448,7 @@
(Logic.list_implies
(mk_distinct qs @
maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop
- (NominalPackage.fresh_const T (fastype_of u) $ t $ u))
+ (Nominal.fresh_const T (fastype_of u) $ t $ u))
args) qs,
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map HOLogic.dest_Trueprop prems))),
@@ -499,13 +499,13 @@
chop (length vc_compat_ths - length args * length qs)
(map (first_order_mrs hyps2) vc_compat_ths);
val vc_compat_ths' =
- NominalPackage.mk_not_sym vc_compat_ths1 @
+ Nominal.mk_not_sym vc_compat_ths1 @
flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2));
val (freshs1, freshs2, ctxt3) = fold
(obtain_fresh_name (args @ map fst qs @ params'))
(map snd qs) ([], [], ctxt2);
- val freshs2' = NominalPackage.mk_not_sym freshs2;
- val pis = map (NominalPackage.perm_of_pair)
+ val freshs2' = Nominal.mk_not_sym freshs2;
+ val pis = map (Nominal.perm_of_pair)
((freshs1 ~~ map fst qs) @ (params' ~~ freshs1));
val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis);
val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
@@ -513,7 +513,7 @@
if x mem args then x
else (case AList.lookup op = tab x of
SOME y => y
- | NONE => fold_rev (NominalPackage.mk_perm []) pis x)
+ | NONE => fold_rev (Nominal.mk_perm []) pis x)
| x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps));
val inst = Thm.first_order_match (Thm.dest_arg
(Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
@@ -522,7 +522,7 @@
rtac (Thm.instantiate inst case_hyp) 1 THEN
SUBPROOF (fn {prems = fresh_hyps, ...} =>
let
- val fresh_hyps' = NominalPackage.mk_not_sym fresh_hyps;
+ val fresh_hyps' = Nominal.mk_not_sym fresh_hyps;
val case_ss = cases_eqvt_ss addsimps freshs2' @
simp_fresh_atm (vc_compat_ths' @ fresh_hyps');
val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh;
@@ -548,13 +548,13 @@
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
val ind_case_names = RuleCases.case_names induct_cases;
- val induct_cases' = InductivePackage.partition_rules' raw_induct
+ val induct_cases' = Inductive.partition_rules' raw_induct
(intrs ~~ induct_cases);
val thss' = map (map atomize_intr) thss;
- val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss');
+ val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
val strong_raw_induct =
- mk_ind_proof ctxt thss' |> InductivePackage.rulify;
- val strong_cases = map (mk_cases_proof ##> InductivePackage.rulify)
+ mk_ind_proof ctxt thss' |> Inductive.rulify;
+ val strong_cases = map (mk_cases_proof ##> Inductive.rulify)
(thsss ~~ elims ~~ cases_prems ~~ cases_prems');
val strong_induct =
if length names > 1 then
@@ -587,17 +587,17 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
val intrs = map atomize_intr intrs;
- val monos = InductivePackage.get_monos ctxt;
- val intrs' = InductivePackage.unpartition_rules intrs
+ val monos = Inductive.get_monos ctxt;
+ val intrs' = Inductive.unpartition_rules intrs
(map (fn (((s, ths), (_, k)), th) =>
- (s, ths ~~ InductivePackage.infer_intro_vars th k ths))
- (InductivePackage.partition_rules raw_induct intrs ~~
- InductivePackage.arities_of raw_induct ~~ elims));
- val k = length (InductivePackage.params_of raw_induct);
+ (s, ths ~~ Inductive.infer_intro_vars th k ths))
+ (Inductive.partition_rules raw_induct intrs ~~
+ Inductive.arities_of raw_induct ~~ elims));
+ val k = length (Inductive.params_of raw_induct);
val atoms' = NominalAtoms.atoms_of thy;
val atoms =
if null xatoms then atoms' else
@@ -635,7 +635,7 @@
val prems'' = map (fn th => Simplifier.simplify eqvt_ss
(mk_perm_bool (cterm_of thy pi) th)) prems';
val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
- map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params)
+ map (cterm_of thy o Nominal.mk_perm [] pi o term_of) params)
intr
in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
end) ctxt' 1 st
@@ -655,7 +655,7 @@
val (ts1, ts2) = chop k ts
in
HOLogic.mk_imp (p, list_comb (h, ts1 @
- map (NominalPackage.mk_perm [] pi') ts2))
+ map (Nominal.mk_perm [] pi') ts2))
end) ps)))
(fn {context, ...} => EVERY (rtac raw_induct 1 :: map (fn intr_vs =>
full_simp_tac eqvt_ss 1 THEN
--- a/src/HOL/Nominal/nominal_inductive2.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML Tue Jun 23 15:48:55 2009 +0100
@@ -56,7 +56,7 @@
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
(Const (s, T), ts) => (case strip_type T of
(Ts, Type (tname, _)) =>
- (case NominalPackage.get_nominal_datatype thy tname of
+ (case Nominal.get_nominal_datatype thy tname of
NONE => fold (add_binders thy i) ts bs
| SOME {descr, index, ...} => (case AList.lookup op =
(#3 (the (AList.lookup op = descr index))) s of
@@ -154,11 +154,11 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
- val ind_params = InductivePackage.params_of raw_induct;
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
+ val ind_params = Inductive.params_of raw_induct;
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
- val monos = InductivePackage.get_monos ctxt;
+ val monos = Inductive.get_monos ctxt;
val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
[] => ()
@@ -249,7 +249,7 @@
| lift_prem t = t;
fun mk_fresh (x, T) = HOLogic.mk_Trueprop
- (NominalPackage.fresh_star_const T fsT $ x $ Bound 0);
+ (Nominal.fresh_star_const T fsT $ x $ Bound 0);
val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) =>
let
@@ -270,7 +270,7 @@
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
HOLogic.list_all (ind_vars, lift_pred p
- (map (fold_rev (NominalPackage.mk_perm ind_Ts)
+ (map (fold_rev (Nominal.mk_perm ind_Ts)
(map Bound (length atomTs downto 1))) ts)))) concls));
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -283,7 +283,7 @@
if null (preds_of ps prem) then SOME prem
else map_term (split_conj (K o I) names) prem prem) prems, q))))
(maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
- (NominalPackage.fresh_star_const U T $ u $ t)) sets)
+ (Nominal.fresh_star_const U T $ u $ t)) sets)
(ts ~~ binder_types (fastype_of p)) @
map (fn (u, U) => HOLogic.mk_Trueprop (Const (@{const_name finite},
HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |>
@@ -339,7 +339,7 @@
val th2' =
Goal.prove ctxt [] []
(list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop
- (f $ fold_rev (NominalPackage.mk_perm (rev pTs))
+ (f $ fold_rev (Nominal.mk_perm (rev pTs))
(pis1 @ pi :: pis2) l $ r)))
(fn _ => cut_facts_tac [th2] 1 THEN
full_simp_tac (HOL_basic_ss addsimps perm_set_forget) 1) |>
@@ -364,7 +364,7 @@
val params' = map term_of cparams'
val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
val pi_sets = map (fn (t, _) =>
- fold_rev (NominalPackage.mk_perm []) pis t) sets';
+ fold_rev (Nominal.mk_perm []) pis t) sets';
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
val gprems1 = List.mapPartial (fn (th, t) =>
if null (preds_of ps t) then SOME th
@@ -380,7 +380,7 @@
in
Goal.prove ctxt' [] []
(HOLogic.mk_Trueprop (list_comb (h,
- map (fold_rev (NominalPackage.mk_perm []) pis) ts)))
+ map (fold_rev (Nominal.mk_perm []) pis) ts)))
(fn _ => simp_tac (HOL_basic_ss addsimps
(fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1)
end) vc_compat_ths vc_compat_vs;
@@ -400,11 +400,11 @@
end;
val pis'' = fold_rev (concat_perm #> map) pis' pis;
val ihyp' = inst_params thy vs_ihypt ihyp
- (map (fold_rev (NominalPackage.mk_perm [])
+ (map (fold_rev (Nominal.mk_perm [])
(pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]);
fun mk_pi th =
Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
- addsimprocs [NominalPackage.perm_simproc])
+ addsimprocs [Nominal.perm_simproc])
(Simplifier.simplify eqvt_ss
(fold_rev (mk_perm_bool o cterm_of thy)
(pis' @ pis) th));
@@ -419,13 +419,13 @@
(fresh_ths2 ~~ sets);
val th = Goal.prove ctxt'' [] []
(HOLogic.mk_Trueprop (list_comb (P $ hd ts,
- map (fold_rev (NominalPackage.mk_perm []) pis') (tl ts))))
+ map (fold_rev (Nominal.mk_perm []) pis') (tl ts))))
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1] @
map (fn th => rtac th 1) fresh_ths3 @
[REPEAT_DETERM_N (length gprems)
(simp_tac (HOL_basic_ss
addsimps [inductive_forall_def']
- addsimprocs [NominalPackage.perm_simproc]) 1 THEN
+ addsimprocs [Nominal.perm_simproc]) 1 THEN
resolve_tac gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
@@ -450,12 +450,12 @@
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
val ind_case_names = RuleCases.case_names induct_cases;
- val induct_cases' = InductivePackage.partition_rules' raw_induct
+ val induct_cases' = Inductive.partition_rules' raw_induct
(intrs ~~ induct_cases);
val thss' = map (map atomize_intr) thss;
- val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss');
+ val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
val strong_raw_induct =
- mk_ind_proof ctxt thss' |> InductivePackage.rulify;
+ mk_ind_proof ctxt thss' |> Inductive.rulify;
val strong_induct =
if length names > 1 then
(strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
--- a/src/HOL/Nominal/nominal_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2095 +0,0 @@
-(* Title: HOL/Nominal/nominal_package.ML
- Author: Stefan Berghofer and Christian Urban, TU Muenchen
-
-Nominal datatype package for Isabelle/HOL.
-*)
-
-signature NOMINAL_PACKAGE =
-sig
- val add_nominal_datatype : DatatypeAux.datatype_config -> string list ->
- (string list * bstring * mixfix *
- (bstring * string list * mixfix) list) list -> theory -> theory
- type descr
- type nominal_datatype_info
- val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
- val get_nominal_datatype : theory -> string -> nominal_datatype_info option
- val mk_perm: typ list -> term -> term -> term
- val perm_of_pair: term * term -> term
- val mk_not_sym: thm list -> thm list
- val perm_simproc: simproc
- val fresh_const: typ -> typ -> term
- val fresh_star_const: typ -> typ -> term
-end
-
-structure NominalPackage : NOMINAL_PACKAGE =
-struct
-
-val finite_emptyI = thm "finite.emptyI";
-val finite_Diff = thm "finite_Diff";
-val finite_Un = thm "finite_Un";
-val Un_iff = thm "Un_iff";
-val In0_eq = thm "In0_eq";
-val In1_eq = thm "In1_eq";
-val In0_not_In1 = thm "In0_not_In1";
-val In1_not_In0 = thm "In1_not_In0";
-val Un_assoc = thm "Un_assoc";
-val Collect_disj_eq = thm "Collect_disj_eq";
-val empty_def = thm "empty_def";
-val empty_iff = thm "empty_iff";
-
-open DatatypeAux;
-open NominalAtoms;
-
-(** FIXME: DatatypePackage should export this function **)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
-
-fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (RuleCases.case_names o exhaust_cases descr o #1)
- (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
-
-end;
-
-(* theory data *)
-
-type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
-
-type nominal_datatype_info =
- {index : int,
- descr : descr,
- sorts : (string * sort) list,
- rec_names : string list,
- rec_rewrites : thm list,
- induction : thm,
- distinct : thm list,
- inject : thm list};
-
-structure NominalDatatypesData = TheoryDataFun
-(
- type T = nominal_datatype_info Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ tabs : T = Symtab.merge (K true) tabs;
-);
-
-val get_nominal_datatypes = NominalDatatypesData.get;
-val put_nominal_datatypes = NominalDatatypesData.put;
-val map_nominal_datatypes = NominalDatatypesData.map;
-val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
-
-
-(**** make datatype info ****)
-
-fun make_dt_info descr sorts induct reccomb_names rec_thms
- (((i, (_, (tname, _, _))), distinct), inject) =
- (tname,
- {index = i,
- descr = descr,
- sorts = sorts,
- rec_names = reccomb_names,
- rec_rewrites = rec_thms,
- induction = induct,
- distinct = distinct,
- inject = inject});
-
-(*******************************)
-
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
-
-
-(** simplification procedure for sorting permutations **)
-
-val dj_cp = thm "dj_cp";
-
-fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
- Type ("fun", [_, U])])) = (T, U);
-
-fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
- | permTs_of _ = [];
-
-fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
- let
- val (aT as Type (a, []), S) = dest_permT T;
- val (bT as Type (b, []), _) = dest_permT U
- in if aT mem permTs_of u andalso aT <> bT then
- let
- val cp = cp_inst_of thy a b;
- val dj = dj_thm_of thy b a;
- val dj_cp' = [cp, dj] MRS dj_cp;
- val cert = SOME o cterm_of thy
- in
- SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
- [cert t, cert r, cert s] dj_cp'))
- end
- else NONE
- end
- | perm_simproc' thy ss _ = NONE;
-
-val perm_simproc =
- Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
-
-val meta_spec = thm "meta_spec";
-
-fun projections rule =
- ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
- |> map (standard #> RuleCases.save rule);
-
-val supp_prod = thm "supp_prod";
-val fresh_prod = thm "fresh_prod";
-val supports_fresh = thm "supports_fresh";
-val supports_def = thm "Nominal.supports_def";
-val fresh_def = thm "fresh_def";
-val supp_def = thm "supp_def";
-val rev_simps = thms "rev.simps";
-val app_simps = thms "append.simps";
-val at_fin_set_supp = thm "at_fin_set_supp";
-val at_fin_set_fresh = thm "at_fin_set_fresh";
-val abs_fun_eq1 = thm "abs_fun_eq1";
-
-val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
-
-fun mk_perm Ts t u =
- let
- val T = fastype_of1 (Ts, t);
- val U = fastype_of1 (Ts, u)
- in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
-
-fun perm_of_pair (x, y) =
- let
- val T = fastype_of x;
- val pT = mk_permT T
- in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
- HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
- end;
-
-fun mk_not_sym ths = maps (fn th => case prop_of th of
- _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
- | _ => [th]) ths;
-
-fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
-fun fresh_star_const T U =
- Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT);
-
-fun gen_add_nominal_datatype prep_typ config new_type_names dts thy =
- let
- (* this theory is used just for parsing *)
-
- val tmp_thy = thy |>
- Theory.copy |>
- Sign.add_types (map (fn (tvs, tname, mx, _) =>
- (Binding.name tname, length tvs, mx)) dts);
-
- val atoms = atoms_of thy;
-
- fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
- let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
- in (constrs @ [(cname, cargs', mx)], sorts') end
-
- fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
- let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
- in (dts @ [(tvs, tname, mx, constrs')], sorts') end
-
- val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
- val tyvars = map (map (fn s =>
- (s, the (AList.lookup (op =) sorts s))) o #1) dts';
-
- fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
- fun augment_sort_typ thy S =
- let val S = Sign.certify_sort thy S
- in map_type_tfree (fn (s, S') => TFree (s,
- if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
- end;
- fun augment_sort thy S = map_types (augment_sort_typ thy S);
-
- val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
- val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
- map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
-
- val ps = map (fn (_, n, _, _) =>
- (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
- val rps = map Library.swap ps;
-
- fun replace_types (Type ("Nominal.ABS", [T, U])) =
- Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
- | replace_types (Type (s, Ts)) =
- Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
- | replace_types T = T;
-
- val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
- map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"),
- map replace_types cargs, NoSyn)) constrs)) dts';
-
- val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
- val full_new_type_names' = map (Sign.full_bname thy) new_type_names';
-
- val ({induction, ...},thy1) =
- DatatypePackage.add_datatype config new_type_names' dts'' thy;
-
- val SOME {descr, ...} = Symtab.lookup
- (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
- fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
-
- val big_name = space_implode "_" new_type_names;
-
-
- (**** define permutation functions ****)
-
- val permT = mk_permT (TFree ("'x", HOLogic.typeS));
- val pi = Free ("pi", permT);
- val perm_types = map (fn (i, _) =>
- let val T = nth_dtyp i
- in permT --> T --> T end) descr;
- val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
- "perm_" ^ name_of_typ (nth_dtyp i)) descr);
- val perm_names = replicate (length new_type_names) "Nominal.perm" @
- map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
- val perm_names_types = perm_names ~~ perm_types;
- val perm_names_types' = perm_names' ~~ perm_types;
-
- val perm_eqs = maps (fn (i, (_, _, constrs)) =>
- let val T = nth_dtyp i
- in map (fn (cname, dts) =>
- let
- val Ts = map (typ_of_dtyp descr sorts) dts;
- val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
- val args = map Free (names ~~ Ts);
- val c = Const (cname, Ts ---> T);
- fun perm_arg (dt, x) =
- let val T = type_of x
- in if is_rec_type dt then
- let val (Us, _) = strip_type T
- in list_abs (map (pair "x") Us,
- Free (nth perm_names_types' (body_index dt)) $ pi $
- list_comb (x, map (fn (i, U) =>
- Const ("Nominal.perm", permT --> U --> U) $
- (Const ("List.rev", permT --> permT) $ pi) $
- Bound i) ((length Us - 1 downto 0) ~~ Us)))
- end
- else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
- end;
- in
- (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Free (nth perm_names_types' i) $
- Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
- list_comb (c, args),
- list_comb (c, map perm_arg (dts ~~ args)))))
- end) constrs
- end) descr;
-
- val (perm_simps, thy2) =
- PrimrecPackage.add_primrec_overloaded
- (map (fn (s, sT) => (s, sT, false))
- (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
- (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
-
- (**** prove that permutation functions introduced by unfolding are ****)
- (**** equivalent to already existing permutation functions ****)
-
- val _ = warning ("length descr: " ^ string_of_int (length descr));
- val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
-
- val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
- val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def";
-
- val unfolded_perm_eq_thms =
- if length descr = length new_type_names then []
- else map standard (List.drop (split_conj_thm
- (Goal.prove_global thy2 [] []
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (c as (s, T), x) =>
- let val [T1, T2] = binder_types T
- in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
- Const ("Nominal.perm", T) $ pi $ Free (x, T2))
- end)
- (perm_names_types ~~ perm_indnames))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac
- (simpset_of thy2 addsimps [perm_fun_def]))])),
- length new_type_names));
-
- (**** prove [] \<bullet> t = t ****)
-
- val _ = warning "perm_empty_thms";
-
- val perm_empty_thms = List.concat (map (fn a =>
- let val permT = mk_permT (Type (a, []))
- in map standard (List.take (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a]
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) => HOLogic.mk_eq
- (Const (s, permT --> T --> T) $
- Const ("List.list.Nil", permT) $ Free (x, T),
- Free (x, T)))
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
- length new_type_names))
- end)
- atoms);
-
- (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
-
- val _ = warning "perm_append_thms";
-
- (*FIXME: these should be looked up statically*)
- val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst";
- val pt2 = PureThy.get_thm thy2 "pt2";
-
- val perm_append_thms = List.concat (map (fn a =>
- let
- val permT = mk_permT (Type (a, []));
- val pi1 = Free ("pi1", permT);
- val pi2 = Free ("pi2", permT);
- val pt_inst = pt_inst_of thy2 a;
- val pt2' = pt_inst RS pt2;
- val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
- in List.take (map standard (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a]
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let val perm = Const (s, permT --> T --> T)
- in HOLogic.mk_eq
- (perm $ (Const ("List.append", permT --> permT --> permT) $
- pi1 $ pi2) $ Free (x, T),
- perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
- end)
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
- length new_type_names)
- end) atoms);
-
- (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
-
- val _ = warning "perm_eq_thms";
-
- val pt3 = PureThy.get_thm thy2 "pt3";
- val pt3_rev = PureThy.get_thm thy2 "pt3_rev";
-
- val perm_eq_thms = List.concat (map (fn a =>
- let
- val permT = mk_permT (Type (a, []));
- val pi1 = Free ("pi1", permT);
- val pi2 = Free ("pi2", permT);
- val at_inst = at_inst_of thy2 a;
- val pt_inst = pt_inst_of thy2 a;
- val pt3' = pt_inst RS pt3;
- val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
- val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
- in List.take (map standard (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
- (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
- permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
- HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let val perm = Const (s, permT --> T --> T)
- in HOLogic.mk_eq
- (perm $ pi1 $ Free (x, T),
- perm $ pi2 $ Free (x, T))
- end)
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames))))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
- length new_type_names)
- end) atoms);
-
- (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
-
- val cp1 = PureThy.get_thm thy2 "cp1";
- val dj_cp = PureThy.get_thm thy2 "dj_cp";
- val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose";
- val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
- val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
-
- fun composition_instance name1 name2 thy =
- let
- val cp_class = cp_class_of thy name1 name2;
- val pt_class =
- if name1 = name2 then [pt_class_of thy name1]
- else [];
- val permT1 = mk_permT (Type (name1, []));
- val permT2 = mk_permT (Type (name2, []));
- val Ts = map body_type perm_types;
- val cp_inst = cp_inst_of thy name1 name2;
- val simps = simpset_of thy addsimps (perm_fun_def ::
- (if name1 <> name2 then
- let val dj = dj_thm_of thy name2 name1
- in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
- else
- let
- val at_inst = at_inst_of thy name1;
- val pt_inst = pt_inst_of thy name1;
- in
- [cp_inst RS cp1 RS sym,
- at_inst RS (pt_inst RS pt_perm_compose) RS sym,
- at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
- end))
- val sort = Sign.certify_sort thy (cp_class :: pt_class);
- val thms = split_conj_thm (Goal.prove_global thy [] []
- (augment_sort thy sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let
- val pi1 = Free ("pi1", permT1);
- val pi2 = Free ("pi2", permT2);
- val perm1 = Const (s, permT1 --> T --> T);
- val perm2 = Const (s, permT2 --> T --> T);
- val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
- in HOLogic.mk_eq
- (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
- perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
- end)
- (perm_names ~~ Ts ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac simps)]))
- in
- fold (fn (s, tvs) => fn thy => AxClass.prove_arity
- (s, map (inter_sort thy sort o snd) tvs, [cp_class])
- (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
- (full_new_type_names' ~~ tyvars) thy
- end;
-
- val (perm_thmss,thy3) = thy2 |>
- fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
- fold (fn atom => fn thy =>
- let val pt_name = pt_class_of thy atom
- in
- fold (fn (s, tvs) => fn thy => AxClass.prove_arity
- (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
- (EVERY
- [Class.intro_classes_tac [],
- resolve_tac perm_empty_thms 1,
- resolve_tac perm_append_thms 1,
- resolve_tac perm_eq_thms 1, assume_tac 1]) thy)
- (full_new_type_names' ~~ tyvars) thy
- end) atoms |>
- PureThy.add_thmss
- [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
- unfolded_perm_eq_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
- perm_empty_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
- perm_append_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
- perm_eq_thms), [Simplifier.simp_add])];
-
- (**** Define representing sets ****)
-
- val _ = warning "representing sets";
-
- val rep_set_names = DatatypeProp.indexify_names
- (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
- val big_rep_name =
- space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
- (fn (i, ("Nominal.noption", _, _)) => NONE
- | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
- val _ = warning ("big_rep_name: " ^ big_rep_name);
-
- fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
- (case AList.lookup op = descr i of
- SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
- apfst (cons dt) (strip_option dt')
- | _ => ([], dtf))
- | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
- apfst (cons dt) (strip_option dt')
- | strip_option dt = ([], dt);
-
- val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
- (List.concat (map (fn (_, (_, _, cs)) => List.concat
- (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
- val dt_atoms = map (fst o dest_Type) dt_atomTs;
-
- fun make_intr s T (cname, cargs) =
- let
- fun mk_prem (dt, (j, j', prems, ts)) =
- let
- val (dts, dt') = strip_option dt;
- val (dts', dt'') = strip_dtyp dt';
- val Ts = map (typ_of_dtyp descr sorts) dts;
- val Us = map (typ_of_dtyp descr sorts) dts';
- val T = typ_of_dtyp descr sorts dt'';
- val free = mk_Free "x" (Us ---> T) j;
- val free' = app_bnds free (length Us);
- fun mk_abs_fun (T, (i, t)) =
- let val U = fastype_of t
- in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
- Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
- end
- in (j + 1, j' + length Ts,
- case dt'' of
- DtRec k => list_all (map (pair "x") Us,
- HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
- T --> HOLogic.boolT) $ free')) :: prems
- | _ => prems,
- snd (List.foldr mk_abs_fun (j', free) Ts) :: ts)
- end;
-
- val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs;
- val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
- list_comb (Const (cname, map fastype_of ts ---> T), ts))
- in Logic.list_implies (prems, concl)
- end;
-
- val (intr_ts, (rep_set_names', recTs')) =
- apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
- (fn ((_, ("Nominal.noption", _, _)), _) => NONE
- | ((i, (_, _, constrs)), rep_set_name) =>
- let val T = nth_dtyp i
- in SOME (map (make_intr rep_set_name T) constrs,
- (rep_set_name, T))
- end)
- (descr ~~ rep_set_names))));
- val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
-
- val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = false, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
- (rep_set_names' ~~ recTs'))
- [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
-
- (**** Prove that representing set is closed under permutation ****)
-
- val _ = warning "proving closure under permutation...";
-
- val abs_perm = PureThy.get_thms thy4 "abs_perm";
-
- val perm_indnames' = List.mapPartial
- (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
- (perm_indnames ~~ descr);
-
- fun mk_perm_closed name = map (fn th => standard (th RS mp))
- (List.take (split_conj_thm (Goal.prove_global thy4 [] []
- (augment_sort thy4
- (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
- (fn ((s, T), x) =>
- let
- val S = Const (s, T --> HOLogic.boolT);
- val permT = mk_permT (Type (name, []))
- in HOLogic.mk_imp (S $ Free (x, T),
- S $ (Const ("Nominal.perm", permT --> T --> T) $
- Free ("pi", permT) $ Free (x, T)))
- end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
- (fn _ => EVERY
- [indtac rep_induct [] 1,
- ALLGOALS (simp_tac (simpset_of thy4 addsimps
- (symmetric perm_fun_def :: abs_perm))),
- ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
- length new_type_names));
-
- val perm_closed_thmss = map mk_perm_closed atoms;
-
- (**** typedef ****)
-
- val _ = warning "defining type...";
-
- val (typedefs, thy6) =
- thy4
- |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
- TypedefPackage.add_typedef false (SOME (Binding.name name'))
- (Binding.name name, map fst tvs, mx)
- (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
- Const (cname, U --> HOLogic.boolT)) NONE
- (rtac exI 1 THEN rtac CollectI 1 THEN
- QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
- let
- val permT = mk_permT
- (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS));
- val pi = Free ("pi", permT);
- val T = Type (Sign.intern_type thy name, map TFree tvs);
- in apfst (pair r o hd)
- (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
- (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
- Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
- (Const ("Nominal.perm", permT --> U --> U) $ pi $
- (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
- Free ("x", T))))), [])] thy)
- end))
- (types_syntax ~~ tyvars ~~
- List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
- new_type_names);
-
- val perm_defs = map snd typedefs;
- val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
- val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
- val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
-
-
- (** prove that new types are in class pt_<name> **)
-
- val _ = warning "prove that new types are in class pt_<name> ...";
-
- fun pt_instance (atom, perm_closed_thms) =
- fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
- perm_def), name), tvs), perm_closed) => fn thy =>
- let
- val pt_class = pt_class_of thy atom;
- val sort = Sign.certify_sort thy
- (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom))
- in AxClass.prove_arity
- (Sign.intern_type thy name,
- map (inter_sort thy sort o snd) tvs, [pt_class])
- (EVERY [Class.intro_classes_tac [],
- rewrite_goals_tac [perm_def],
- asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
- asm_full_simp_tac (simpset_of thy addsimps
- [Rep RS perm_closed RS Abs_inverse]) 1,
- asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
- ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
- end)
- (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
- new_type_names ~~ tyvars ~~ perm_closed_thms);
-
-
- (** prove that new types are in class cp_<name1>_<name2> **)
-
- val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
-
- fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
- let
- val cp_class = cp_class_of thy atom1 atom2;
- val sort = Sign.certify_sort thy
- (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
- (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
- pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2)));
- val cp1' = cp_inst_of thy atom1 atom2 RS cp1
- in fold (fn ((((((Abs_inverse, Rep),
- perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
- AxClass.prove_arity
- (Sign.intern_type thy name,
- map (inter_sort thy sort o snd) tvs, [cp_class])
- (EVERY [Class.intro_classes_tac [],
- rewrite_goals_tac [perm_def],
- asm_full_simp_tac (simpset_of thy addsimps
- ((Rep RS perm_closed1 RS Abs_inverse) ::
- (if atom1 = atom2 then []
- else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
- cong_tac 1,
- rtac refl 1,
- rtac cp1' 1]) thy)
- (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
- tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
- end;
-
- val thy7 = fold (fn x => fn thy => thy |>
- pt_instance x |>
- fold (cp_instance x) (atoms ~~ perm_closed_thmss))
- (atoms ~~ perm_closed_thmss) thy6;
-
- (**** constructors ****)
-
- fun mk_abs_fun (x, t) =
- let
- val T = fastype_of x;
- val U = fastype_of t
- in
- Const ("Nominal.abs_fun", T --> U --> T -->
- Type ("Nominal.noption", [U])) $ x $ t
- end;
-
- val (ty_idxs, _) = List.foldl
- (fn ((i, ("Nominal.noption", _, _)), p) => p
- | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
-
- fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
- | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
- | reindex dt = dt;
-
- fun strip_suffix i s = implode (List.take (explode s, size s - i));
-
- (** strips the "_Rep" in type names *)
- fun strip_nth_name i s =
- let val xs = Long_Name.explode s;
- in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
-
- val (descr'', ndescr) = ListPair.unzip (map_filter
- (fn (i, ("Nominal.noption", _, _)) => NONE
- | (i, (s, dts, constrs)) =>
- let
- val SOME index = AList.lookup op = ty_idxs i;
- val (constrs2, constrs1) =
- map_split (fn (cname, cargs) =>
- apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
- (fold_map (fn dt => fn dts =>
- let val (dts', dt') = strip_option dt
- in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
- cargs [])) constrs
- in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)),
- (index, constrs2))
- end) descr);
-
- val (descr1, descr2) = chop (length new_type_names) descr'';
- val descr' = [descr1, descr2];
-
- fun partition_cargs idxs xs = map (fn (i, j) =>
- (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
-
- val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
- map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
- (constrs ~~ idxss)))) (descr'' ~~ ndescr);
-
- fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
-
- val rep_names = map (fn s =>
- Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
- val abs_names = map (fn s =>
- Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
-
- val recTs = get_rec_types descr'' sorts;
- val newTs' = Library.take (length new_type_names, recTs');
- val newTs = Library.take (length new_type_names, recTs);
-
- val full_new_type_names = map (Sign.full_bname thy) new_type_names;
-
- fun make_constr_def tname T T' ((thy, defs, eqns),
- (((cname_rep, _), (cname, cargs)), (cname', mx))) =
- let
- fun constr_arg ((dts, dt), (j, l_args, r_args)) =
- let
- val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
- (dts ~~ (j upto j + length dts - 1))
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ x :: l_args,
- List.foldr mk_abs_fun
- (case dt of
- DtRec k => if k < length new_type_names then
- Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
- typ_of_dtyp descr sorts dt) $ x
- else error "nested recursion not (yet) supported"
- | _ => x) xs :: r_args)
- end
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
- val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
- val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
- val constrT = map fastype_of l_args ---> T;
- val lhs = list_comb (Const (cname, constrT), l_args);
- val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
- val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const (rep_name, T --> T') $ lhs, rhs));
- val def_name = (Long_Name.base_name cname) ^ "_def";
- val ([def_thm], thy') = thy |>
- Sign.add_consts_i [(Binding.name cname', constrT, mx)] |>
- (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
- in (thy', defs @ [def_thm], eqns @ [eqn]) end;
-
- fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
- (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
- let
- val rep_const = cterm_of thy
- (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
- val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
- ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
- in
- (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
- end;
-
- val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
- List.take (pdescr, length new_type_names) ~~
- new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
-
- val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
- val rep_inject_thms = map (#Rep_inject o fst) typedefs
-
- (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *)
-
- fun prove_constr_rep_thm eqn =
- let
- val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
- val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
- in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
- [resolve_tac inj_thms 1,
- rewrite_goals_tac rewrites,
- rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac Rep_thms 1)])
- end;
-
- val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
-
- (* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
-
- fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
- let
- val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
- val Type ("fun", [T, U]) = fastype_of Rep;
- val permT = mk_permT (Type (atom, []));
- val pi = Free ("pi", permT);
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8
- (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
- Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
- (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
- perm_closed_thms @ Rep_thms)) 1)
- end) Rep_thms;
-
- val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
- (atoms ~~ perm_closed_thmss));
-
- (* prove distinctness theorems *)
-
- val distinct_props = DatatypeProp.make_distincts descr' sorts;
- val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
- dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
- constr_rep_thmss dist_lemmas;
-
- fun prove_distinct_thms _ (_, []) = []
- | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) =
- let
- val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
- simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
- in dist_thm :: standard (dist_thm RS not_sym) ::
- prove_distinct_thms p (k, ts)
- end;
-
- val distinct_thms = map2 prove_distinct_thms
- (constr_rep_thmss ~~ dist_lemmas) distinct_props;
-
- (** prove equations for permutation functions **)
-
- val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
- let val T = nth_dtyp' i
- in List.concat (map (fn (atom, perm_closed_thms) =>
- map (fn ((cname, dts), constr_rep_thm) =>
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
- val permT = mk_permT (Type (atom, []));
- val pi = Free ("pi", permT);
-
- fun perm t =
- let val T = fastype_of t
- in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
-
- fun constr_arg ((dts, dt), (j, l_args, r_args)) =
- let
- val Ts = map (typ_of_dtyp descr'' sorts) dts;
- val xs = map (fn (T, i) => mk_Free "x" T i)
- (Ts ~~ (j upto j + length dts - 1))
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ x :: l_args,
- map perm (xs @ [x]) @ r_args)
- end
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts;
- val c = Const (cname, map fastype_of l_args ---> T)
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8
- (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
- (fn _ => EVERY
- [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
- simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
- constr_defs @ perm_closed_thms)) 1,
- TRY (simp_tac (HOL_basic_ss addsimps
- (symmetric perm_fun_def :: abs_perm)) 1),
- TRY (simp_tac (HOL_basic_ss addsimps
- (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
- perm_closed_thms)) 1)])
- end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
-
- (** prove injectivity of constructors **)
-
- val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
- val alpha = PureThy.get_thms thy8 "alpha";
- val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
-
- val pt_cp_sort =
- map (pt_class_of thy8) dt_atoms @
- maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
-
- val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
- let val T = nth_dtyp' i
- in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
- if null dts then NONE else SOME
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
-
- fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
- let
- val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
- val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
- val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ (x :: args1), ys @ (y :: args2),
- HOLogic.mk_eq
- (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs)
- end;
-
- val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts;
- val Ts = map fastype_of args1;
- val c = Const (cname, Ts ---> T)
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
- foldr1 HOLogic.mk_conj eqs))))
- (fn _ => EVERY
- [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
- rep_inject_thms')) 1,
- TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
- alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
- perm_rep_perm_thms)) 1)])
- end) (constrs ~~ constr_rep_thms)
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
-
- (** equations for support and freshness **)
-
- val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
- (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
- let val T = nth_dtyp' i
- in List.concat (map (fn (cname, dts) => map (fn atom =>
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
- val atomT = Type (atom, []);
-
- fun process_constr ((dts, dt), (j, args1, args2)) =
- let
- val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2)
- end;
-
- val (_, args1, args2) = List.foldr process_constr (1, [], []) dts;
- val Ts = map fastype_of args1;
- val c = list_comb (Const (cname, Ts ---> T), args1);
- fun supp t =
- Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
- fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
- val supp_thm = Goal.prove_global thy8 [] []
- (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (supp c,
- if null dts then HOLogic.mk_set atomT []
- else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2)))))
- (fn _ =>
- simp_tac (HOL_basic_ss addsimps (supp_def ::
- Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
- symmetric empty_def :: finite_emptyI :: simp_thms @
- abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
- in
- (supp_thm,
- Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fresh c,
- if null dts then HOLogic.true_const
- else foldr1 HOLogic.mk_conj (map fresh args2)))))
- (fn _ =>
- simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
- end) atoms) constrs)
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
-
- (**** weak induction theorem ****)
-
- fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
- let
- val Rep_t = Const (List.nth (rep_names, i), T --> U) $
- mk_Free "x" T i;
-
- val Abs_t = Const (List.nth (abs_names, i), U --> T)
-
- in (prems @ [HOLogic.imp $
- (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
- (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
- concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
- end;
-
- val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
-
- val indrule_lemma = Goal.prove_global thy8 [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
- [REPEAT (etac conjE 1),
- REPEAT (EVERY
- [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
- etac mp 1, resolve_tac Rep_thms 1])]);
-
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
- val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
- map (Free o apfst fst o dest_Var) Ps;
- val indrule_lemma' = cterm_instantiate
- (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
-
- val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
-
- val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
- val dt_induct = Goal.prove_global thy8 []
- (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
- (fn {prems, ...} => EVERY
- [rtac indrule_lemma' 1,
- (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms' 1),
- simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
- DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
- (prems ~~ constr_defs))]);
-
- val case_names_induct = mk_case_names_induct descr'';
-
- (**** prove that new datatypes have finite support ****)
-
- val _ = warning "proving finite support for the new datatype";
-
- val indnames = DatatypeProp.make_tnames recTs;
-
- val abs_supp = PureThy.get_thms thy8 "abs_supp";
- val supp_atm = PureThy.get_thms thy8 "supp_atm";
-
- val finite_supp_thms = map (fn atom =>
- let val atomT = Type (atom, [])
- in map standard (List.take
- (split_conj_thm (Goal.prove_global thy8 [] []
- (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
- (HOLogic.mk_Trueprop
- (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
- Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
- (indnames ~~ recTs)))))
- (fn _ => indtac dt_induct indnames 1 THEN
- ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
- (abs_supp @ supp_atm @
- PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
- List.concat supp_thms))))),
- length new_type_names))
- end) atoms;
-
- val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
-
- (* Function to add both the simp and eqvt attributes *)
- (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
-
- val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
-
- val (_, thy9) = thy8 |>
- Sign.add_path big_name |>
- PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
- PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
- Sign.parent_path ||>>
- DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
- DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
- DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
- DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
- DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
- DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
- fold (fn (atom, ths) => fn thy =>
- let
- val class = fs_class_of thy atom;
- val sort = Sign.certify_sort thy (class :: pt_cp_sort)
- in fold (fn Type (s, Ts) => AxClass.prove_arity
- (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
- (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
- end) (atoms ~~ finite_supp_thms);
-
- (**** strong induction theorem ****)
-
- val pnames = if length descr'' = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
- val ind_sort = if null dt_atomTs then HOLogic.typeS
- else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
- val fsT = TFree ("'n", ind_sort);
- val fsT' = TFree ("'n", HOLogic.typeS);
-
- val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
- (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
-
- fun make_pred fsT i T =
- Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
-
- fun mk_fresh1 xs [] = []
- | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
- (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
- (filter (fn (_, U) => T = U) (rev xs)) @
- mk_fresh1 (y :: xs) ys;
-
- fun mk_fresh2 xss [] = []
- | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
- map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
- (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @
- mk_fresh2 (p :: xss) yss;
-
- fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
- let
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr'' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr'' sorts) recs;
- val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = tnames ~~ Ts;
- val frees' = partition_cargs idxs frees;
- val z = (Name.variant tnames "z", fsT);
-
- fun mk_prem ((dt, s), T) =
- let
- val (Us, U) = strip_type T;
- val l = length Us
- in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
- end;
-
- val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
- val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
- (f T (Free p) (Free z))) (List.concat (map fst frees')) @
- mk_fresh1 [] (List.concat (map fst frees')) @
- mk_fresh2 [] frees'
-
- in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
- HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
- list_comb (Const (cname, Ts ---> T), map Free frees))))
- end;
-
- val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
- map (make_ind_prem fsT (fn T => fn t => fn u =>
- fresh_const T fsT $ t $ u) i T)
- (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
- val tnames = DatatypeProp.make_tnames recTs;
- val zs = Name.variant_list tnames (replicate (length descr'') "z");
- val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), tname), z) =>
- make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
- (descr'' ~~ recTs ~~ tnames ~~ zs)));
- val induct = Logic.list_implies (ind_prems, ind_concl);
-
- val ind_prems' =
- map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
- HOLogic.mk_Trueprop (Const ("Finite_Set.finite",
- (snd (split_last (binder_types T)) --> HOLogic.boolT) -->
- HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
- List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
- map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
- HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
- (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
- val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), tname), z) =>
- make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
- (descr'' ~~ recTs ~~ tnames ~~ zs)));
- val induct' = Logic.list_implies (ind_prems', ind_concl');
-
- val aux_ind_vars =
- (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
- map mk_permT dt_atomTs) @ [("z", fsT')];
- val aux_ind_Ts = rev (map snd aux_ind_vars);
- val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn (((i, _), T), tname) =>
- HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
- fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
- (Free (tname, T))))
- (descr'' ~~ recTs ~~ tnames)));
-
- val fin_set_supp = map (fn s =>
- at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
- val fin_set_fresh = map (fn s =>
- at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
- val pt1_atoms = map (fn Type (s, _) =>
- PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
- val pt2_atoms = map (fn Type (s, _) =>
- PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
- val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
- val fs_atoms = PureThy.get_thms thy9 "fin_supp";
- val abs_supp = PureThy.get_thms thy9 "abs_supp";
- val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh";
- val calc_atm = PureThy.get_thms thy9 "calc_atm";
- val fresh_atm = PureThy.get_thms thy9 "fresh_atm";
- val fresh_left = PureThy.get_thms thy9 "fresh_left";
- val perm_swap = PureThy.get_thms thy9 "perm_swap";
-
- fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
- let
- val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
- val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
- (HOLogic.exists_const T $ Abs ("x", T,
- fresh_const T (fastype_of p) $
- Bound 0 $ p)))
- (fn _ => EVERY
- [resolve_tac exists_fresh' 1,
- simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
- fin_set_supp @ ths)) 1]);
- val (([cx], ths), ctxt') = Obtain.result
- (fn _ => EVERY
- [etac exE 1,
- full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
- REPEAT (etac conjE 1)])
- [ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
-
- fun fresh_fresh_inst thy a b =
- let
- val T = fastype_of a;
- val SOME th = find_first (fn th => case prop_of th of
- _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
- | _ => false) perm_fresh_fresh
- in
- Drule.instantiate' []
- [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
- end;
-
- val fs_cp_sort =
- map (fs_class_of thy9) dt_atoms @
- maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
-
- (**********************************************************************
- The subgoals occurring in the proof of induct_aux have the
- following parameters:
-
- x_1 ... x_k p_1 ... p_m z
-
- where
-
- x_i : constructor arguments (introduced by weak induction rule)
- p_i : permutations (one for each atom type in the data type)
- z : freshness context
- ***********************************************************************)
-
- val _ = warning "proving strong induction theorem ...";
-
- val induct_aux = Goal.prove_global thy9 []
- (map (augment_sort thy9 fs_cp_sort) ind_prems')
- (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
- let
- val (prems1, prems2) = chop (length dt_atomTs) prems;
- val ind_ss2 = HOL_ss addsimps
- finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
- val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
- fresh_atm @ rev_simps @ app_simps;
- val ind_ss3 = HOL_ss addsimps abs_fun_eq1 ::
- abs_perm @ calc_atm @ perm_swap;
- val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
- fin_set_fresh @ calc_atm;
- val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
- val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
- val th = Goal.prove context [] []
- (augment_sort thy9 fs_cp_sort aux_ind_concl)
- (fn {context = context1, ...} =>
- EVERY (indtac dt_induct tnames 1 ::
- maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
- map (fn ((cname, cargs), is) =>
- REPEAT (rtac allI 1) THEN
- SUBPROOF (fn {prems = iprems, params, concl,
- context = context2, ...} =>
- let
- val concl' = term_of concl;
- val _ $ (_ $ _ $ u) = concl';
- val U = fastype_of u;
- val (xs, params') =
- chop (length cargs) (map term_of params);
- val Ts = map fastype_of xs;
- val cnstr = Const (cname, Ts ---> U);
- val (pis, z) = split_last params';
- val mk_pi = fold_rev (mk_perm []) pis;
- val xs' = partition_cargs is xs;
- val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
- val ts = maps (fn (ts, u) => ts @ [u]) xs'';
- val (freshs1, freshs2, context3) = fold (fn t =>
- let val T = fastype_of t
- in obtain_fresh_name' prems1
- (the (AList.lookup op = fresh_fs T) $ z :: ts) T
- end) (maps fst xs') ([], [], context2);
- val freshs1' = unflat (map fst xs') freshs1;
- val freshs2' = map (Simplifier.simplify ind_ss4)
- (mk_not_sym freshs2);
- val ind_ss1' = ind_ss1 addsimps freshs2';
- val ind_ss3' = ind_ss3 addsimps freshs2';
- val rename_eq =
- if forall (null o fst) xs' then []
- else [Goal.prove context3 [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (list_comb (cnstr, ts),
- list_comb (cnstr, maps (fn ((bs, t), cs) =>
- cs @ [fold_rev (mk_perm []) (map perm_of_pair
- (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
- (fn _ => EVERY
- (simp_tac (HOL_ss addsimps flat inject_thms) 1 ::
- REPEAT (FIRSTGOAL (rtac conjI)) ::
- maps (fn ((bs, t), cs) =>
- if null bs then []
- else rtac sym 1 :: maps (fn (b, c) =>
- [rtac trans 1, rtac sym 1,
- rtac (fresh_fresh_inst thy9 b c) 1,
- simp_tac ind_ss1' 1,
- simp_tac ind_ss2 1,
- simp_tac ind_ss3' 1]) (bs ~~ cs))
- (xs'' ~~ freshs1')))];
- val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
- [simp_tac (ind_ss6 addsimps rename_eq) 1,
- cut_facts_tac iprems 1,
- (resolve_tac prems THEN_ALL_NEW
- SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
- _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
- simp_tac ind_ss1' i
- | _ $ (Const ("Not", _) $ _) =>
- resolve_tac freshs2' i
- | _ => asm_simp_tac (HOL_basic_ss addsimps
- pt2_atoms addsimprocs [perm_simproc]) i)) 1])
- val final = ProofContext.export context3 context2 [th]
- in
- resolve_tac final 1
- end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
- in
- EVERY
- [cut_facts_tac [th] 1,
- REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
- REPEAT (etac allE 1),
- REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
- end);
-
- val induct_aux' = Thm.instantiate ([],
- map (fn (s, v as Var (_, T)) =>
- (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
- (pnames ~~ map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
- map (fn (_, f) =>
- let val f' = Logic.varify f
- in (cterm_of thy9 f',
- cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
- end) fresh_fs) induct_aux;
-
- val induct = Goal.prove_global thy9 []
- (map (augment_sort thy9 fs_cp_sort) ind_prems)
- (augment_sort thy9 fs_cp_sort ind_concl)
- (fn {prems, ...} => EVERY
- [rtac induct_aux' 1,
- REPEAT (resolve_tac fs_atoms 1),
- REPEAT ((resolve_tac prems THEN_ALL_NEW
- (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
-
- val (_, thy10) = thy9 |>
- Sign.add_path big_name |>
- PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
- PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
- PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
-
- (**** recursion combinator ****)
-
- val _ = warning "defining recursion combinator ...";
-
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
-
- val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
-
- val rec_sort = if null dt_atomTs then HOLogic.typeS else
- Sign.certify_sort thy10 pt_cp_sort;
-
- val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
- val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
-
- val rec_set_Ts = map (fn (T1, T2) =>
- rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
-
- val big_rec_name = big_name ^ "_rec_set";
- val rec_set_names' =
- if length descr'' = 1 then [big_rec_name] else
- map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
- (1 upto (length descr''));
- val rec_set_names = map (Sign.full_bname thy10) rec_set_names';
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
- val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
- (rec_set_names' ~~ rec_set_Ts);
- val rec_sets = map (fn c => list_comb (Const c, rec_fns))
- (rec_set_names ~~ rec_set_Ts);
-
- (* introduction rules for graph of recursion function *)
-
- val rec_preds = map (fn (a, T) =>
- Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
-
- fun mk_fresh3 rs [] = []
- | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
- List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
- else SOME (HOLogic.mk_Trueprop
- (fresh_const T U $ Free y $ Free r))) rs) ys) @
- mk_fresh3 rs yss;
-
- (* FIXME: avoid collisions with other variable names? *)
- val rec_ctxt = Free ("z", fsT');
-
- fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
- rec_eq_prems, l), ((cname, cargs), idxs)) =
- let
- val Ts = map (typ_of_dtyp descr'' sorts) cargs;
- val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
- val frees' = partition_cargs idxs frees;
- val binders = List.concat (map fst frees');
- val atomTs = distinct op = (maps (map snd o fst) frees');
- val recs = List.mapPartial
- (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
- (partition_cargs idxs cargs ~~ frees');
- val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
- map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
- val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
- (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
- val prems2 =
- map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
- (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
- val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
- val prems4 = map (fn ((i, _), y) =>
- HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
- val prems5 = mk_fresh3 (recs ~~ frees'') frees';
- val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
- frees'') atomTs;
- val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
- (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
- val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
- val result_freshs = map (fn p as (_, T) =>
- fresh_const T (fastype_of result) $ Free p $ result) binders;
- val P = HOLogic.mk_Trueprop (p $ result)
- in
- (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
- HOLogic.mk_Trueprop (rec_set $
- list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
- rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
- rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
- Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
- HOLogic.mk_Trueprop fr))) result_freshs,
- rec_eq_prems @ [List.concat prems2 @ prems3],
- l + 1)
- end;
-
- val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
- Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
- Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
- (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
-
- val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
- thy10 |>
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
- (map dest_Free rec_fns)
- (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
- PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct");
-
- (** equivariance **)
-
- val fresh_bij = PureThy.get_thms thy11 "fresh_bij";
- val perm_bij = PureThy.get_thms thy11 "perm_bij";
-
- val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
- let
- val permT = mk_permT aT;
- val pi = Free ("pi", permT);
- val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
- (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
- val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
- (rec_set_names ~~ rec_set_Ts);
- val ps = map (fn ((((T, U), R), R'), i) =>
- let
- val x = Free ("x" ^ string_of_int i, T);
- val y = Free ("y" ^ string_of_int i, U)
- in
- (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
- val ths = map (fn th => standard (th RS mp)) (split_conj_thm
- (Goal.prove_global thy11 [] []
- (augment_sort thy1 pt_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
- (fn _ => rtac rec_induct 1 THEN REPEAT
- (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
- addsimps flat perm_simps'
- addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- (resolve_tac rec_intrs THEN_ALL_NEW
- asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
- val ths' = map (fn ((P, Q), th) =>
- Goal.prove_global thy11 [] []
- (augment_sort thy1 pt_cp_sort
- (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
- (fn _ => dtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), permT)),
- cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
- NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
- in (ths, ths') end) dt_atomTs);
-
- (** finite support **)
-
- val rec_fin_supp_thms = map (fn aT =>
- let
- val name = Long_Name.base_name (fst (dest_Type aT));
- val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
- val aset = HOLogic.mk_setT aT;
- val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
- val fins = map (fn (f, T) => HOLogic.mk_Trueprop
- (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
- (rec_fns ~~ rec_fn_Ts)
- in
- map (fn th => standard (th RS mp)) (split_conj_thm
- (Goal.prove_global thy11 []
- (map (augment_sort thy11 fs_cp_sort) fins)
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (((T, U), R), i) =>
- let
- val x = Free ("x" ^ string_of_int i, T);
- val y = Free ("y" ^ string_of_int i, U)
- in
- HOLogic.mk_imp (R $ x $ y,
- finite $ (Const ("Nominal.supp", U --> aset) $ y))
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
- (1 upto length recTs))))))
- (fn {prems = fins, ...} =>
- (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
- (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
- end) dt_atomTs;
-
- (** freshness **)
-
- val finite_premss = map (fn aT =>
- map (fn (f, T) => HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
- (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
-
- val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
-
- val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
- let
- val name = Long_Name.base_name (fst (dest_Type aT));
- val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
- val a = Free ("a", aT);
- val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
- (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
- in
- map (fn (((T, U), R), eqvt_th) =>
- let
- val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
- val y = Free ("y", U);
- val y' = Free ("y'", U)
- in
- standard (Goal.prove (ProofContext.init thy11) []
- (map (augment_sort thy11 fs_cp_sort)
- (finite_prems @
- [HOLogic.mk_Trueprop (R $ x $ y),
- HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
- HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
- HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
- freshs))
- (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
- (fn {prems, context} =>
- let
- val (finite_prems, rec_prem :: unique_prem ::
- fresh_prems) = chop (length finite_prems) prems;
- val unique_prem' = unique_prem RS spec RS mp;
- val unique = [unique_prem', unique_prem' RS sym] MRS trans;
- val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
- val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
- in EVERY
- [rtac (Drule.cterm_instantiate
- [(cterm_of thy11 S,
- cterm_of thy11 (Const ("Nominal.supp",
- fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
- supports_fresh) 1,
- simp_tac (HOL_basic_ss addsimps
- [supports_def, symmetric fresh_def, fresh_prod]) 1,
- REPEAT_DETERM (resolve_tac [allI, impI] 1),
- REPEAT_DETERM (etac conjE 1),
- rtac unique 1,
- SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
- [cut_facts_tac [rec_prem] 1,
- rtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
- cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
- asm_simp_tac (HOL_ss addsimps
- (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
- rtac rec_prem 1,
- simp_tac (HOL_ss addsimps (fs_name ::
- supp_prod :: finite_Un :: finite_prems)) 1,
- simp_tac (HOL_ss addsimps (symmetric fresh_def ::
- fresh_prod :: fresh_prems)) 1]
- end))
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
- end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
-
- (** uniqueness **)
-
- val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
- val fun_tupleT = fastype_of fun_tuple;
- val rec_unique_frees =
- DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
- val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
- val rec_unique_frees' =
- DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
- val rec_unique_concls = map (fn ((x, U), R) =>
- Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
- Abs ("y", U, R $ Free x $ Bound 0))
- (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
-
- val induct_aux_rec = Drule.cterm_instantiate
- (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
- (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
- Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
- fresh_fs @
- map (fn (((P, T), (x, U)), Q) =>
- (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
- Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
- (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
- map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
- rec_unique_frees)) induct_aux;
-
- fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
- let
- val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
- val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
- (HOLogic.exists_const T $ Abs ("x", T,
- fresh_const T (fastype_of p) $ Bound 0 $ p)))
- (fn _ => EVERY
- [cut_facts_tac ths 1,
- REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
- resolve_tac exists_fresh' 1,
- asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
- val (([cx], ths), ctxt') = Obtain.result
- (fn _ => EVERY
- [etac exE 1,
- full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
- REPEAT (etac conjE 1)])
- [ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
-
- val finite_ctxt_prems = map (fn aT =>
- HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
-
- val rec_unique_thms = split_conj_thm (Goal.prove
- (ProofContext.init thy11) (map fst rec_unique_frees)
- (map (augment_sort thy11 fs_cp_sort)
- (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
- (fn {prems, context} =>
- let
- val k = length rec_fns;
- val (finite_thss, ths1) = fold_map (fn T => fn xs =>
- apfst (pair T) (chop k xs)) dt_atomTs prems;
- val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
- val (P_ind_ths, fcbs) = chop k ths2;
- val P_ths = map (fn th => th RS mp) (split_conj_thm
- (Goal.prove context
- (map fst (rec_unique_frees'' @ rec_unique_frees')) []
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (((x, y), S), P) => HOLogic.mk_imp
- (S $ Free x $ Free y, P $ (Free y)))
- (rec_unique_frees'' ~~ rec_unique_frees' ~~
- rec_sets ~~ rec_preds)))))
- (fn _ =>
- rtac rec_induct 1 THEN
- REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
- val rec_fin_supp_thms' = map
- (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
- (rec_fin_supp_thms ~~ finite_thss);
- in EVERY
- ([rtac induct_aux_rec 1] @
- maps (fn ((_, finite_ths), finite_th) =>
- [cut_facts_tac (finite_th :: finite_ths) 1,
- asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
- (finite_thss ~~ finite_ctxt_ths) @
- maps (fn ((_, idxss), elim) => maps (fn idxs =>
- [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
- REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
- rtac ex1I 1,
- (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
- rotate_tac ~1 1,
- ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
- (HOL_ss addsimps List.concat distinct_thms)) 1] @
- (if null idxs then [] else [hyp_subst_tac 1,
- SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
- let
- val SOME prem = find_first (can (HOLogic.dest_eq o
- HOLogic.dest_Trueprop o prop_of)) prems';
- val _ $ (_ $ lhs $ rhs) = prop_of prem;
- val _ $ (_ $ lhs' $ rhs') = term_of concl;
- val rT = fastype_of lhs';
- val (c, cargsl) = strip_comb lhs;
- val cargsl' = partition_cargs idxs cargsl;
- val boundsl = List.concat (map fst cargsl');
- val (_, cargsr) = strip_comb rhs;
- val cargsr' = partition_cargs idxs cargsr;
- val boundsr = List.concat (map fst cargsr');
- val (params1, _ :: params2) =
- chop (length params div 2) (map term_of params);
- val params' = params1 @ params2;
- val rec_prems = filter (fn th => case prop_of th of
- _ $ p => (case head_of p of
- Const (s, _) => s mem rec_set_names
- | _ => false)
- | _ => false) prems';
- val fresh_prems = filter (fn th => case prop_of th of
- _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
- | _ $ (Const ("Not", _) $ _) => true
- | _ => false) prems';
- val Ts = map fastype_of boundsl;
-
- val _ = warning "step 1: obtaining fresh names";
- val (freshs1, freshs2, context'') = fold
- (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
- (List.concat (map snd finite_thss) @
- finite_ctxt_ths @ rec_prems)
- rec_fin_supp_thms')
- Ts ([], [], context');
- val pi1 = map perm_of_pair (boundsl ~~ freshs1);
- val rpi1 = rev pi1;
- val pi2 = map perm_of_pair (boundsr ~~ freshs1);
- val rpi2 = rev pi2;
-
- val fresh_prems' = mk_not_sym fresh_prems;
- val freshs2' = mk_not_sym freshs2;
-
- (** as, bs, cs # K as ts, K bs us **)
- val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
- val prove_fresh_ss = HOL_ss addsimps
- (finite_Diff :: List.concat fresh_thms @
- fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
- (* FIXME: avoid asm_full_simp_tac ? *)
- fun prove_fresh ths y x = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const
- (fastype_of x) (fastype_of y) $ x $ y))
- (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
- val constr_fresh_thms =
- map (prove_fresh fresh_prems lhs) boundsl @
- map (prove_fresh fresh_prems rhs) boundsr @
- map (prove_fresh freshs2 lhs) freshs1 @
- map (prove_fresh freshs2 rhs) freshs1;
-
- (** pi1 o (K as ts) = pi2 o (K bs us) **)
- val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
- val pi1_pi2_eq = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
- (fn _ => EVERY
- [cut_facts_tac constr_fresh_thms 1,
- asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
- rtac prem 1]);
-
- (** pi1 o ts = pi2 o us **)
- val _ = warning "step 4: pi1 o ts = pi2 o us";
- val pi1_pi2_eqs = map (fn (t, u) =>
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
- (fn _ => EVERY
- [cut_facts_tac [pi1_pi2_eq] 1,
- asm_full_simp_tac (HOL_ss addsimps
- (calc_atm @ List.concat perm_simps' @
- fresh_prems' @ freshs2' @ abs_perm @
- alpha @ List.concat inject_thms)) 1]))
- (map snd cargsl' ~~ map snd cargsr');
-
- (** pi1^-1 o pi2 o us = ts **)
- val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
- val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
- (fn _ => simp_tac (HOL_ss addsimps
- ((eq RS sym) :: perm_swap)) 1))
- (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
-
- val (rec_prems1, rec_prems2) =
- chop (length rec_prems div 2) rec_prems;
-
- (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
- val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
- val rec_prems' = map (fn th =>
- let
- val _ $ (S $ x $ y) = prop_of th;
- val Const (s, _) = head_of S;
- val k = find_index (equal s) rec_set_names;
- val pi = rpi1 @ pi2;
- fun mk_pi z = fold_rev (mk_perm []) pi z;
- fun eqvt_tac p =
- let
- val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
- val l = find_index (equal T) dt_atomTs;
- val th = List.nth (List.nth (rec_equiv_thms', l), k);
- val th' = Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), U)),
- cterm_of thy11 p)]) th;
- in rtac th' 1 end;
- val th' = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
- (fn _ => EVERY
- (map eqvt_tac pi @
- [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
- perm_swap @ perm_fresh_fresh)) 1,
- rtac th 1]))
- in
- Simplifier.simplify
- (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
- end) rec_prems2;
-
- val ihs = filter (fn th => case prop_of th of
- _ $ (Const ("All", _) $ _) => true | _ => false) prems';
-
- (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
- val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
- val rec_eqns = map (fn (th, ih) =>
- let
- val th' = th RS (ih RS spec RS mp) RS sym;
- val _ $ (_ $ lhs $ rhs) = prop_of th';
- fun strip_perm (_ $ _ $ t) = strip_perm t
- | strip_perm t = t;
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 lhs,
- fold_rev (mk_perm []) pi2 (strip_perm rhs))))
- (fn _ => simp_tac (HOL_basic_ss addsimps
- (th' :: perm_swap)) 1)
- end) (rec_prems' ~~ ihs);
-
- (** as # rs **)
- val _ = warning "step 8: as # rs";
- val rec_freshs = List.concat
- (map (fn (rec_prem, ih) =>
- let
- val _ $ (S $ x $ (y as Free (_, T))) =
- prop_of rec_prem;
- val k = find_index (equal S) rec_sets;
- val atoms = List.concat (List.mapPartial (fn (bs, z) =>
- if z = x then NONE else SOME bs) cargsl')
- in
- map (fn a as Free (_, aT) =>
- let val l = find_index (equal aT) dt_atomTs;
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
- (fn _ => EVERY
- (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
- map (fn th => rtac th 1)
- (snd (List.nth (finite_thss, l))) @
- [rtac rec_prem 1, rtac ih 1,
- REPEAT_DETERM (resolve_tac fresh_prems 1)]))
- end) atoms
- end) (rec_prems1 ~~ ihs));
-
- (** as # fK as ts rs , bs # fK bs us vs **)
- val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
- fun prove_fresh_result (a as Free (_, aT)) =
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
- (fn _ => EVERY
- [resolve_tac fcbs 1,
- REPEAT_DETERM (resolve_tac
- (fresh_prems @ rec_freshs) 1),
- REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
- THEN resolve_tac rec_prems 1),
- resolve_tac P_ind_ths 1,
- REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
-
- val fresh_results'' = map prove_fresh_result boundsl;
-
- fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
- let val th' = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $
- fold_rev (mk_perm []) (rpi2 @ pi1) a $
- fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
- (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
- rtac th 1)
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
- (fn _ => EVERY
- [cut_facts_tac [th'] 1,
- full_simp_tac (Simplifier.theory_context thy11 HOL_ss
- addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
- addsimprocs [NominalPermeq.perm_simproc_app]) 1,
- full_simp_tac (HOL_ss addsimps (calc_atm @
- fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
- end;
-
- val fresh_results = fresh_results'' @ map prove_fresh_result''
- (boundsl ~~ boundsr ~~ fresh_results'');
-
- (** cs # fK as ts rs , cs # fK bs us vs **)
- val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
- fun prove_fresh_result' recs t (a as Free (_, aT)) =
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
- (fn _ => EVERY
- [cut_facts_tac recs 1,
- REPEAT_DETERM (dresolve_tac
- (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
- NominalPermeq.fresh_guess_tac
- (HOL_ss addsimps (freshs2 @
- fs_atoms @ fresh_atm @
- List.concat (map snd finite_thss))) 1]);
-
- val fresh_results' =
- map (prove_fresh_result' rec_prems1 rhs') freshs1 @
- map (prove_fresh_result' rec_prems2 lhs') freshs1;
-
- (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
- val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
- val pi1_pi2_result = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
- (fn _ => simp_tac (Simplifier.context context'' HOL_ss
- addsimps pi1_pi2_eqs @ rec_eqns
- addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- TRY (simp_tac (HOL_ss addsimps
- (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
-
- val _ = warning "final result";
- val final = Goal.prove context'' [] [] (term_of concl)
- (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
- full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
- fresh_results @ fresh_results') 1);
- val final' = ProofContext.export context'' context' [final];
- val _ = warning "finished!"
- in
- resolve_tac final' 1
- end) context 1])) idxss) (ndescr ~~ rec_elims))
- end));
-
- val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
-
- (* define primrec combinators *)
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.full_bname thy11)
- (if length descr'' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr''))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- val (reccomb_defs, thy12) =
- thy11
- |> Sign.add_consts_i (map (fn ((name, T), T') =>
- (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
- (reccomb_names ~~ recTs ~~ rec_result_Ts))
- |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
- (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
- Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
- set $ Free ("x", T) $ Free ("y", T'))))))
- (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
-
- (* prove characteristic equations for primrec combinators *)
-
- val rec_thms = map (fn (prems, concl) =>
- let
- val _ $ (_ $ (_ $ x) $ _) = concl;
- val (_, cargs) = strip_comb x;
- val ps = map (fn (x as Free (_, T), i) =>
- (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
- val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
- val prems' = List.concat finite_premss @ finite_ctxt_prems @
- rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
- fun solve rules prems = resolve_tac rules THEN_ALL_NEW
- (resolve_tac prems THEN_ALL_NEW atac)
- in
- Goal.prove_global thy12 []
- (map (augment_sort thy12 fs_cp_sort) prems')
- (augment_sort thy12 fs_cp_sort concl')
- (fn {prems, ...} => EVERY
- [rewrite_goals_tac reccomb_defs,
- rtac the1_equality 1,
- solve rec_unique_thms prems 1,
- resolve_tac rec_intrs 1,
- REPEAT (solve (prems @ rec_total_thms) prems 1)])
- end) (rec_eq_prems ~~
- DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
-
- val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
- ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
-
- (* FIXME: theorems are stored in database for testing only *)
- val (_, thy13) = thy12 |>
- PureThy.add_thmss
- [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
- ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
- ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
- ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
- ((Binding.name "rec_unique", map standard rec_unique_thms), []),
- ((Binding.name "recs", rec_thms), [])] ||>
- Sign.parent_path ||>
- map_nominal_datatypes (fold Symtab.update dt_infos);
-
- in
- thy13
- end;
-
-val add_nominal_datatype = gen_add_nominal_datatype DatatypePackage.read_typ;
-
-
-(* FIXME: The following stuff should be exported by DatatypePackage *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
- (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
-
-fun mk_datatype args =
- let
- val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
- val specs = map (fn ((((_, vs), t), mx), cons) =>
- (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
- in add_nominal_datatype DatatypeAux.default_datatype_config names specs end;
-
-val _ =
- OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-end;
-
-end
--- a/src/HOL/Nominal/nominal_primrec.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML Tue Jun 23 15:48:55 2009 +0100
@@ -3,7 +3,7 @@
Author: Stefan Berghofer, TU Muenchen
Package for defining functions on nominal datatypes by primitive recursion.
-Taken from HOL/Tools/primrec_package.ML
+Taken from HOL/Tools/primrec.ML
*)
signature NOMINAL_PRIMREC =
@@ -223,7 +223,7 @@
(* find datatypes which contain all datatypes in tnames' *)
-fun find_dts (dt_info : NominalPackage.nominal_datatype_info Symtab.table) _ [] = []
+fun find_dts (dt_info : Nominal.nominal_datatype_info Symtab.table) _ [] = []
| find_dts dt_info tnames' (tname::tnames) =
(case Symtab.lookup dt_info tname of
NONE => primrec_err (quote tname ^ " is not a nominal datatype")
@@ -247,7 +247,7 @@
val eqns' = map unquantify spec'
val eqns = fold_rev (process_eqn lthy (fn v => Variable.is_fixed lthy v
orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) spec' [];
- val dt_info = NominalPackage.get_nominal_datatypes (ProofContext.theory_of lthy);
+ val dt_info = Nominal.get_nominal_datatypes (ProofContext.theory_of lthy);
val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) =>
map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns
val _ =
--- a/src/HOL/Parity.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Parity.thy Tue Jun 23 15:48:55 2009 +0100
@@ -35,7 +35,7 @@
lemma even_zero_nat[simp]: "even (0::nat)" by presburger
-lemma odd_zero_nat [simp]: "odd (1::nat)" by presburger
+lemma odd_1_nat [simp]: "odd (1::nat)" by presburger
declare even_def[of "number_of v", standard, simp]
--- a/src/HOL/Product_Type.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Product_Type.thy Tue Jun 23 15:48:55 2009 +0100
@@ -9,9 +9,9 @@
imports Inductive
uses
("Tools/split_rule.ML")
- ("Tools/inductive_set_package.ML")
+ ("Tools/inductive_set.ML")
("Tools/inductive_realizer.ML")
- ("Tools/datatype_package/datatype_realizer.ML")
+ ("Tools/Datatype/datatype_realizer.ML")
begin
subsection {* @{typ bool} is a datatype *}
@@ -399,7 +399,7 @@
by (simp add: split_def id_def)
lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f"
- -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
+ -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity Datatype. *}
by (rule ext) auto
lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)"
@@ -734,7 +734,7 @@
text {*
@{term prod_fun} --- action of the product functor upon
- functions.
+ Datatypes.
*}
definition prod_fun :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
@@ -1151,10 +1151,10 @@
use "Tools/inductive_realizer.ML"
setup InductiveRealizer.setup
-use "Tools/inductive_set_package.ML"
-setup InductiveSetPackage.setup
+use "Tools/inductive_set.ML"
+setup Inductive_Set.setup
-use "Tools/datatype_package/datatype_realizer.ML"
+use "Tools/Datatype/datatype_realizer.ML"
setup DatatypeRealizer.setup
end
--- a/src/HOL/Rational.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Rational.thy Tue Jun 23 15:48:55 2009 +0100
@@ -851,11 +851,11 @@
subsection {* Implementation of rational numbers as pairs of integers *}
-lemma Fract_norm: "Fract (a div zgcd a b) (b div zgcd a b) = Fract a b"
+lemma Fract_norm: "Fract (a div gcd a b) (b div gcd a b) = Fract a b"
proof (cases "a = 0 \<or> b = 0")
case True then show ?thesis by (auto simp add: eq_rat)
next
- let ?c = "zgcd a b"
+ let ?c = "gcd a b"
case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
then have "?c \<noteq> 0" by simp
then have "Fract ?c ?c = Fract 1 1" by (simp add: eq_rat)
@@ -870,7 +870,7 @@
definition Fract_norm :: "int \<Rightarrow> int \<Rightarrow> rat" where
[simp, code del]: "Fract_norm a b = Fract a b"
-lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
+lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = gcd a b in
if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))"
by (simp add: eq_rat Zero_rat_def Let_def Fract_norm)
--- a/src/HOL/RealDef.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/RealDef.thy Tue Jun 23 15:48:55 2009 +0100
@@ -895,14 +895,15 @@
lemma Rats_abs_nat_div_natE:
assumes "x \<in> \<rat>"
- obtains m n where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
+ obtains m n :: nat
+ where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
proof -
from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
by(auto simp add: Rats_eq_int_div_nat)
hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
let ?gcd = "gcd m n"
- from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by (simp add: gcd_zero)
+ from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
let ?k = "m div ?gcd"
let ?l = "n div ?gcd"
let ?gcd' = "gcd ?k ?l"
@@ -924,9 +925,9 @@
have "?gcd' = 1"
proof -
have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
- by (rule gcd_mult_distrib2)
+ by (rule nat_gcd_mult_distrib)
with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
- with gcd show ?thesis by simp
+ with gcd show ?thesis by auto
qed
ultimately show ?thesis ..
qed
--- a/src/HOL/Recdef.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Recdef.thy Tue Jun 23 15:48:55 2009 +0100
@@ -16,7 +16,7 @@
("Tools/TFL/thry.ML")
("Tools/TFL/tfl.ML")
("Tools/TFL/post.ML")
- ("Tools/recdef_package.ML")
+ ("Tools/recdef.ML")
begin
text{** This form avoids giant explosions in proofs. NOTE USE OF ==*}
@@ -76,8 +76,8 @@
use "Tools/TFL/thry.ML"
use "Tools/TFL/tfl.ML"
use "Tools/TFL/post.ML"
-use "Tools/recdef_package.ML"
-setup RecdefPackage.setup
+use "Tools/recdef.ML"
+setup Recdef.setup
lemmas [recdef_simp] =
inv_image_def
--- a/src/HOL/Record.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Record.thy Tue Jun 23 15:48:55 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Record.thy
- ID: $Id$
Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
*)
@@ -7,7 +6,7 @@
theory Record
imports Product_Type
-uses ("Tools/record_package.ML")
+uses ("Tools/record.ML")
begin
lemma prop_subst: "s = t \<Longrightarrow> PROP P t \<Longrightarrow> PROP P s"
@@ -56,7 +55,7 @@
"_record_scheme" :: "[fields, 'a] => 'a" ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
"_record_update" :: "['a, updates] => 'b" ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
-use "Tools/record_package.ML"
-setup RecordPackage.setup
+use "Tools/record.ML"
+setup Record.setup
end
--- a/src/HOL/Statespace/state_fun.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Statespace/state_fun.ML Tue Jun 23 15:48:55 2009 +0100
@@ -74,7 +74,7 @@
val string_eq_simp_tac =
simp_tac (HOL_basic_ss
addsimps (thms "list.inject"@thms "char.inject"@simp_thms)
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc]
addcongs [thm "block_conj_cong"])
end;
@@ -89,7 +89,7 @@
in
val lookup_ss = (HOL_basic_ss
addsimps (thms "list.inject"@thms "char.inject"@simp_thms@rules)
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc]
addcongs [thm "block_conj_cong"]
addSolver StateSpace.distinctNameSolver)
end;
@@ -167,7 +167,7 @@
val meta_ext = thm "StateFun.meta_ext";
val o_apply = thm "Fun.o_apply";
val ss' = (HOL_ss addsimps (update_apply::o_apply::thms "list.inject"@thms "char.inject")
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc,StateSpace.distinct_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc,StateSpace.distinct_simproc]
addcongs [thm "block_conj_cong"]);
in
val update_simproc =
@@ -267,7 +267,7 @@
val swap_ex_eq = thm "StateFun.swap_ex_eq";
fun is_selector thy T sel =
let
- val (flds,more) = RecordPackage.get_recT_fields thy T
+ val (flds,more) = Record.get_recT_fields thy T
in member (fn (s,(n,_)) => n=s) (more::flds) sel
end;
in
@@ -340,7 +340,7 @@
| mkName (TFree (x,_)) = mkUpper (Long_Name.base_name x)
| mkName (TVar ((x,_),_)) = mkUpper (Long_Name.base_name x);
-fun is_datatype thy n = is_some (Symtab.lookup (DatatypePackage.get_datatypes thy) n);
+fun is_datatype thy n = is_some (Symtab.lookup (Datatype.get_datatypes thy) n);
fun mk_map "List.list" = Syntax.const "List.map"
| mk_map n = Syntax.const ("StateFun.map_" ^ Long_Name.base_name n);
--- a/src/HOL/Statespace/state_space.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Statespace/state_space.ML Tue Jun 23 15:48:55 2009 +0100
@@ -585,8 +585,8 @@
end
handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name);
-val define_statespace = gen_define_statespace RecordPackage.read_typ NONE;
-val define_statespace_i = gen_define_statespace RecordPackage.cert_typ;
+val define_statespace = gen_define_statespace Record.read_typ NONE;
+val define_statespace_i = gen_define_statespace Record.cert_typ;
(*** parse/print - translations ***)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,704 @@
+(* Title: HOL/Tools/datatype.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package for Isabelle/HOL.
+*)
+
+signature DATATYPE =
+sig
+ include DATATYPE_COMMON
+ type rules = {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list}
+ val add_datatype : config -> string list -> (string list * binding * mixfix *
+ (binding * typ list * mixfix) list) list -> theory -> rules * theory
+ val datatype_cmd : string list -> (string list * binding * mixfix *
+ (binding * string list * mixfix) list) list -> theory -> theory
+ val rep_datatype : config -> (rules -> Proof.context -> Proof.context)
+ -> string list option -> term list -> theory -> Proof.state
+ val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
+ val get_datatypes : theory -> info Symtab.table
+ val get_datatype : theory -> string -> info option
+ val the_datatype : theory -> string -> info
+ val datatype_of_constr : theory -> string -> info option
+ val datatype_of_case : theory -> string -> info option
+ val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
+ val the_datatype_descr : theory -> string list
+ -> descr * (string * sort) list * string list
+ * (string list * string list) * (typ list * typ list)
+ val get_datatype_constrs : theory -> string -> (string * typ) list option
+ val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
+ val distinct_simproc : simproc
+ val make_case : Proof.context -> bool -> string list -> term ->
+ (term * term) list -> term * (term * (int * bool)) list
+ val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
+ val read_typ: theory ->
+ (typ list * (string * sort) list) * string -> typ list * (string * sort) list
+ val setup: theory -> theory
+end;
+
+structure Datatype : DATATYPE =
+struct
+
+open DatatypeAux;
+
+
+(* theory data *)
+
+structure DatatypesData = TheoryDataFun
+(
+ type T =
+ {types: info Symtab.table,
+ constrs: info Symtab.table,
+ cases: info Symtab.table};
+
+ val empty =
+ {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ({types = types1, constrs = constrs1, cases = cases1},
+ {types = types2, constrs = constrs2, cases = cases2}) =
+ {types = Symtab.merge (K true) (types1, types2),
+ constrs = Symtab.merge (K true) (constrs1, constrs2),
+ cases = Symtab.merge (K true) (cases1, cases2)};
+);
+
+val get_datatypes = #types o DatatypesData.get;
+val map_datatypes = DatatypesData.map;
+
+
+(** theory information about datatypes **)
+
+fun put_dt_infos (dt_infos : (string * info) list) =
+ map_datatypes (fn {types, constrs, cases} =>
+ {types = fold Symtab.update dt_infos types,
+ constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
+ (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
+ (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
+ cases = fold Symtab.update
+ (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
+ cases});
+
+val get_datatype = Symtab.lookup o get_datatypes;
+
+fun the_datatype thy name = (case get_datatype thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
+
+val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
+val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
+
+fun get_datatype_descr thy dtco =
+ get_datatype thy dtco
+ |> Option.map (fn info as { descr, index, ... } =>
+ (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
+
+fun the_datatype_spec thy dtco =
+ let
+ val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
+ val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
+ o DatatypeAux.dest_DtTFree) dtys;
+ val cos = map
+ (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
+ in (sorts, cos) end;
+
+fun the_datatype_descr thy (raw_tycos as raw_tyco :: _) =
+ let
+ val info = the_datatype thy raw_tyco;
+ val descr = #descr info;
+
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
+ val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
+ o dest_DtTFree) dtys;
+
+ fun is_DtTFree (DtTFree _) = true
+ | is_DtTFree _ = false
+ val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
+ val protoTs as (dataTs, _) = chop k descr
+ |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
+
+ val tycos = map fst dataTs;
+ val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
+ else error ("Type constructors " ^ commas (map quote raw_tycos)
+ ^ "do not belong exhaustively to one mutual recursive datatype");
+
+ val (Ts, Us) = (pairself o map) Type protoTs;
+
+ val names = map Long_Name.base_name (the_default tycos (#alt_names info));
+ val (auxnames, _) = Name.make_context names
+ |> fold_map (yield_singleton Name.variants o name_of_typ) Us
+
+ in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
+
+fun get_datatype_constrs thy dtco =
+ case try (the_datatype_spec thy) dtco
+ of SOME (sorts, cos) =>
+ let
+ fun subst (v, sort) = TVar ((v, 0), sort);
+ fun subst_ty (TFree v) = subst v
+ | subst_ty ty = ty;
+ val dty = Type (dtco, map subst sorts);
+ fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
+ in SOME (map mk_co cos) end
+ | NONE => NONE;
+
+
+(** induct method setup **)
+
+(* case names *)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = maps dt_recs dts
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct (op =) (maps dt_recs args));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
+
+fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (RuleCases.case_names o exhaust_cases descr o #1)
+ (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
+
+end;
+
+fun add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs cong_att =
+ PureThy.add_thmss [((Binding.name "simps", simps), []),
+ ((Binding.empty, flat case_thms @
+ flat distinct @ rec_thms), [Simplifier.simp_add]),
+ ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
+ ((Binding.empty, flat inject), [iff_add]),
+ ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
+ ((Binding.empty, weak_case_congs), [cong_att])]
+ #> snd;
+
+
+(* add_cases_induct *)
+
+fun add_cases_induct infos induction thy =
+ let
+ val inducts = ProjectRule.projections (ProofContext.init thy) induction;
+
+ fun named_rules (name, {index, exhaustion, ...}: info) =
+ [((Binding.empty, nth inducts index), [Induct.induct_type name]),
+ ((Binding.empty, exhaustion), [Induct.cases_type name])];
+ fun unnamed_rule i =
+ ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
+ in
+ thy |> PureThy.add_thms
+ (maps named_rules infos @
+ map unnamed_rule (length infos upto length inducts - 1)) |> snd
+ |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
+ end;
+
+
+
+(**** simplification procedure for showing distinctness of constructors ****)
+
+fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
+ | stripT p = p;
+
+fun stripC (i, f $ x) = stripC (i + 1, f)
+ | stripC p = p;
+
+val distinctN = "constr_distinct";
+
+fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
+ FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
+ atac 2, resolve_tac thms 1, etac FalseE 1]))
+ | ManyConstrs (thm, simpset) =>
+ let
+ val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
+ map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
+ ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
+ in
+ Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
+ full_simp_tac (Simplifier.inherit_context ss simpset) 1,
+ REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
+ eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
+ etac FalseE 1]))
+ end;
+
+fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
+ (case (stripC (0, t1), stripC (0, t2)) of
+ ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
+ (case (stripT (0, T1), stripT (0, T2)) of
+ ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
+ if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
+ (case (get_datatype_descr thy) tname1 of
+ SOME (_, (_, constrs)) => let val cnames = map fst constrs
+ in if cname1 mem cnames andalso cname2 mem cnames then
+ SOME (distinct_rule thy ss tname1
+ (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
+ else NONE
+ end
+ | NONE => NONE)
+ else NONE
+ | _ => NONE)
+ | _ => NONE)
+ | distinct_proc _ _ _ = NONE;
+
+val distinct_simproc =
+ Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
+
+val dist_ss = HOL_ss addsimprocs [distinct_simproc];
+
+val simproc_setup =
+ Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
+
+
+(**** translation rules for case ****)
+
+fun make_case ctxt = DatatypeCase.make_case
+ (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
+
+fun strip_case ctxt = DatatypeCase.strip_case
+ (datatype_of_case (ProofContext.theory_of ctxt));
+
+fun add_case_tr' case_names thy =
+ Sign.add_advanced_trfuns ([], [],
+ map (fn case_name =>
+ let val case_name' = Sign.const_syntax_name thy case_name
+ in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
+ end) case_names, []) thy;
+
+val trfun_setup =
+ Sign.add_advanced_trfuns ([],
+ [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
+ [], []);
+
+
+(* prepare types *)
+
+fun read_typ thy ((Ts, sorts), str) =
+ let
+ val ctxt = ProofContext.init thy
+ |> fold (Variable.declare_typ o TFree) sorts;
+ val T = Syntax.read_typ ctxt str;
+ in (Ts @ [T], Term.add_tfreesT T sorts) end;
+
+fun cert_typ sign ((Ts, sorts), raw_T) =
+ let
+ val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
+ TYPE (msg, _, _) => error msg;
+ val sorts' = Term.add_tfreesT T sorts;
+ in (Ts @ [T],
+ case duplicates (op =) (map fst sorts') of
+ [] => sorts'
+ | dups => error ("Inconsistent sort constraints for " ^ commas dups))
+ end;
+
+
+(**** make datatype info ****)
+
+fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
+ (((((((((i, (_, (tname, _, _))), case_name), case_thms),
+ exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
+ (tname,
+ {index = i,
+ alt_names = alt_names,
+ descr = descr,
+ sorts = sorts,
+ rec_names = reccomb_names,
+ rec_rewrites = rec_thms,
+ case_name = case_name,
+ case_rewrites = case_thms,
+ induction = induct,
+ exhaustion = exhaustion_thm,
+ distinct = distinct_thm,
+ inject = inject,
+ nchotomy = nchotomy,
+ case_cong = case_cong,
+ weak_case_cong = weak_case_cong});
+
+type rules = {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list}
+
+structure DatatypeInterpretation = InterpretationFun
+ (type T = config * string list val eq: T * T -> bool = eq_snd op =);
+fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
+
+
+(******************* definitional introduction of datatypes *******************)
+
+fun add_datatype_def (config : config) new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy =
+ let
+ val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
+ DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts
+ types_syntax constr_syntax case_names_induct;
+
+ val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names descr
+ sorts induct case_names_exhausts thy2;
+ val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
+ config new_type_names descr sorts dt_info inject dist_rewrites
+ (Simplifier.theory_context thy3 dist_ss) induct thy3;
+ val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
+ config new_type_names descr sorts reccomb_names rec_thms thy4;
+ val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms config new_type_names
+ descr sorts inject dist_rewrites casedist_thms case_thms thy6;
+ val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys config new_type_names
+ descr sorts casedist_thms thy7;
+ val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
+ descr sorts nchotomys case_thms thy8;
+ val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ descr sorts thy9;
+
+ val dt_infos = map
+ (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
+ ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
+ casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy12 =
+ thy10
+ |> add_case_tr' case_names
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
+ |> DatatypeInterpretation.data (config, map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct,
+ simps = simps}, thy12)
+ end;
+
+
+(*********************** declare existing type as datatype *********************)
+
+fun prove_rep_datatype (config : config) alt_names new_type_names descr sorts induct inject half_distinct thy =
+ let
+ val ((_, [induct']), _) =
+ Variable.importT_thms [induct] (Variable.thm_context induct);
+
+ fun err t = error ("Ill-formed predicate in induction rule: " ^
+ Syntax.string_of_term_global thy t);
+
+ fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
+ ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
+ | get_typ t = err t;
+ val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
+
+ val dt_info = get_datatypes thy;
+
+ val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
+ val (case_names_induct, case_names_exhausts) =
+ (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
+
+ val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val (casedist_thms, thy2) = thy |>
+ DatatypeAbsProofs.prove_casedist_thms config new_type_names [descr] sorts induct
+ case_names_exhausts;
+ val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
+ config new_type_names [descr] sorts dt_info inject distinct
+ (Simplifier.theory_context thy2 dist_ss) induct thy2;
+ val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms
+ config new_type_names [descr] sorts reccomb_names rec_thms thy3;
+ val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
+ config new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
+ val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys config new_type_names
+ [descr] sorts casedist_thms thy5;
+ val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
+ [descr] sorts nchotomys case_thms thy6;
+ val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ [descr] sorts thy7;
+
+ val ((_, [induct']), thy10) =
+ thy8
+ |> store_thmss "inject" new_type_names inject
+ ||>> store_thmss "distinct" new_type_names distinct
+ ||> Sign.add_path (space_implode "_" new_type_names)
+ ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
+
+ val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
+ ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
+ map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy11 =
+ thy10
+ |> add_case_tr' case_names
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct'
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
+ |> snd
+ |> DatatypeInterpretation.data (config, map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct',
+ simps = simps}, thy11)
+ end;
+
+fun gen_rep_datatype prep_term (config : config) after_qed alt_names raw_ts thy =
+ let
+ fun constr_of_term (Const (c, T)) = (c, T)
+ | constr_of_term t =
+ error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+ fun no_constr (c, T) = error ("Bad constructor: "
+ ^ Sign.extern_const thy c ^ "::"
+ ^ Syntax.string_of_typ_global thy T);
+ fun type_of_constr (cT as (_, T)) =
+ let
+ val frees = OldTerm.typ_tfrees T;
+ val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
+ handle TYPE _ => no_constr cT
+ val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
+ val _ = if length frees <> length vs then no_constr cT else ();
+ in (tyco, (vs, cT)) end;
+
+ val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
+ val _ = case map_filter (fn (tyco, _) =>
+ if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
+ of [] => ()
+ | tycos => error ("Type(s) " ^ commas (map quote tycos)
+ ^ " already represented inductivly");
+ val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
+ val ms = case distinct (op =) (map length raw_vss)
+ of [n] => 0 upto n - 1
+ | _ => error ("Different types in given constructors");
+ fun inter_sort m = map (fn xs => nth xs m) raw_vss
+ |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
+ val sorts = map inter_sort ms;
+ val vs = Name.names Name.context Name.aT sorts;
+
+ fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
+ (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
+
+ val cs = map (apsnd (map norm_constr)) raw_cs;
+ val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
+ o fst o strip_type;
+ val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
+
+ fun mk_spec (i, (tyco, constr)) = (i, (tyco,
+ map (DtTFree o fst) vs,
+ (map o apsnd) dtyps_of_typ constr))
+ val descr = map_index mk_spec cs;
+ val injs = DatatypeProp.make_injs [descr] vs;
+ val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
+ val ind = DatatypeProp.make_ind [descr] vs;
+ val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
+
+ fun after_qed' raw_thms =
+ let
+ val [[[induct]], injs, half_distincts] =
+ unflat rules (map Drule.zero_var_indexes_list raw_thms);
+ (*FIXME somehow dubious*)
+ in
+ ProofContext.theory_result
+ (prove_rep_datatype config alt_names new_type_names descr vs induct injs half_distincts)
+ #-> after_qed
+ end;
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+ end;
+
+val rep_datatype = gen_rep_datatype Sign.cert_term;
+val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I);
+
+
+
+(******************************** add datatype ********************************)
+
+fun gen_add_datatype prep_typ (config : config) new_type_names dts thy =
+ let
+ val _ = Theory.requires thy "Datatype" "datatype definitions";
+
+ (* this theory is used just for parsing *)
+
+ val tmp_thy = thy |>
+ Theory.copy |>
+ Sign.add_types (map (fn (tvs, tname, mx, _) =>
+ (tname, length tvs, mx)) dts);
+
+ val (tyvars, _, _, _)::_ = dts;
+ val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
+ let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
+ in (case duplicates (op =) tvs of
+ [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
+ else error ("Mutually recursive datatypes must have same type parameters")
+ | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
+ " : " ^ commas dups))
+ end) dts);
+
+ val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
+ [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
+
+ fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
+ let
+ fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
+ let
+ val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
+ val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
+ [] => ()
+ | vs => error ("Extra type variables on rhs: " ^ commas vs))
+ in (constrs @ [((if #flat_names config then Sign.full_name tmp_thy else
+ Sign.full_name_path tmp_thy tname')
+ (Binding.map_name (Syntax.const_name mx') cname),
+ map (dtyp_of_typ new_dts) cargs')],
+ constr_syntax' @ [(cname, mx')], sorts'')
+ end handle ERROR msg => cat_error msg
+ ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
+ " of datatype " ^ quote (Binding.str_of tname));
+
+ val (constrs', constr_syntax', sorts') =
+ fold prep_constr constrs ([], [], sorts)
+
+ in
+ case duplicates (op =) (map fst constrs') of
+ [] =>
+ (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
+ map DtTFree tvs, constrs'))],
+ constr_syntax @ [constr_syntax'], sorts', i + 1)
+ | dups => error ("Duplicate constructors " ^ commas dups ^
+ " in datatype " ^ quote (Binding.str_of tname))
+ end;
+
+ val (dts', constr_syntax, sorts', i) =
+ fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
+ val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
+ val dt_info = get_datatypes thy;
+ val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
+ val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
+ if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
+ else raise exn;
+
+ val descr' = flat descr;
+ val case_names_induct = mk_case_names_induct descr';
+ val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
+ in
+ add_datatype_def
+ (config : config) new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy
+ end;
+
+val add_datatype = gen_add_datatype cert_typ;
+val datatype_cmd = snd ooo gen_add_datatype read_typ default_config;
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ DatatypeRepProofs.distinctness_limit_setup #>
+ simproc_setup #>
+ trfun_setup #>
+ DatatypeInterpretation.init;
+
+
+(* outer syntax *)
+
+local
+
+structure P = OuterParse and K = OuterKeyword
+
+fun prep_datatype_decls args =
+ let
+ val names = map
+ (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
+ val specs = map (fn ((((_, vs), t), mx), cons) =>
+ (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
+ in (names, specs) end;
+
+val parse_datatype_decl =
+ (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
+ (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix)));
+
+val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls;
+
+in
+
+val _ =
+ OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
+ (parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs)));
+
+val _ =
+ OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
+ (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
+ >> (fn (alt_names, ts) => Toplevel.print
+ o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+
+end;
+
+
+(* document antiquotation *)
+
+val _ = ThyOutput.antiquotation "datatype" Args.tyname
+ (fn {source = src, context = ctxt, ...} => fn dtco =>
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val (vs, cos) = the_datatype_spec thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
+ Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
+ | pretty_typ_bracket ty =
+ Syntax.pretty_typ ctxt ty;
+ fun pretty_constr (co, tys) =
+ (Pretty.block o Pretty.breaks)
+ (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
+ map pretty_typ_bracket tys);
+ val pretty_datatype =
+ Pretty.block
+ (Pretty.command "datatype" :: Pretty.brk 1 ::
+ Syntax.pretty_typ ctxt ty ::
+ Pretty.str " =" :: Pretty.brk 1 ::
+ flat (separate [Pretty.brk 1, Pretty.str "| "]
+ (map (single o pretty_constr) cos)));
+ in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
+
+end;
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,447 @@
+(* Title: HOL/Tools/datatype_abs_proofs.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Proofs and defintions independent of concrete representation
+of datatypes (i.e. requiring only abstract properties such as
+injectivity / distinctness of constructors and induction)
+
+ - case distinction (exhaustion) theorems
+ - characteristic equations for primrec combinators
+ - characteristic equations for case combinators
+ - equations for splitting "P (case ...)" expressions
+ - "nchotomy" and "case_cong" theorems for TFL
+*)
+
+signature DATATYPE_ABS_PROOFS =
+sig
+ include DATATYPE_COMMON
+ val prove_casedist_thms : config -> string list ->
+ descr list -> (string * sort) list -> thm ->
+ attribute list -> theory -> thm list * theory
+ val prove_primrec_thms : config -> string list ->
+ descr list -> (string * sort) list ->
+ info Symtab.table -> thm list list -> thm list list ->
+ simpset -> thm -> theory -> (string list * thm list) * theory
+ val prove_case_thms : config -> string list ->
+ descr list -> (string * sort) list ->
+ string list -> thm list -> theory -> (thm list list * string list) * theory
+ val prove_split_thms : config -> string list ->
+ descr list -> (string * sort) list ->
+ thm list list -> thm list list -> thm list -> thm list list -> theory ->
+ (thm * thm) list * theory
+ val prove_nchotomys : config -> string list -> descr list ->
+ (string * sort) list -> thm list -> theory -> thm list * theory
+ val prove_weak_case_congs : string list -> descr list ->
+ (string * sort) list -> theory -> thm list * theory
+ val prove_case_congs : string list ->
+ descr list -> (string * sort) list ->
+ thm list -> thm list list -> theory -> thm list * theory
+end;
+
+structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS =
+struct
+
+open DatatypeAux;
+
+(************************ case distinction theorems ***************************)
+
+fun prove_casedist_thms (config : config) new_type_names descr sorts induct case_names_exhausts thy =
+ let
+ val _ = message config "Proving case distinction theorems ...";
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ val {maxidx, ...} = rep_thm induct;
+ val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+ fun prove_casedist_thm ((i, t), T) =
+ let
+ val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
+ Abs ("z", T', Const ("True", T''))) induct_Ps;
+ val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
+ Var (("P", 0), HOLogic.boolT))
+ val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs)));
+ val cert = cterm_of thy;
+ val insts' = (map cert induct_Ps) ~~ (map cert insts);
+ val induct' = refl RS ((List.nth
+ (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
+
+ in
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} => EVERY
+ [rtac induct' 1,
+ REPEAT (rtac TrueI 1),
+ REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
+ REPEAT (rtac TrueI 1)])
+ end;
+
+ val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
+ (DatatypeProp.make_casedists descr sorts) ~~ newTs)
+ in
+ thy
+ |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms
+ end;
+
+
+(*************************** primrec combinators ******************************)
+
+fun prove_primrec_thms (config : config) new_type_names descr sorts
+ (dt_info : info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
+ let
+ val _ = message config "Constructing primrec combinators ...";
+
+ val big_name = space_implode "_" new_type_names;
+ val thy0 = add_path (#flat_names config) big_name thy;
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+ val big_rec_name' = big_name ^ "_rec_set";
+ val rec_set_names' =
+ if length descr' = 1 then [big_rec_name'] else
+ map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
+ (1 upto (length descr'));
+ val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
+
+ val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
+
+ val rec_set_Ts = map (fn (T1, T2) =>
+ reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
+ val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
+ (rec_set_names' ~~ rec_set_Ts);
+ val rec_sets = map (fn c => list_comb (Const c, rec_fns))
+ (rec_set_names ~~ rec_set_Ts);
+
+ (* introduction rules for graph of primrec function *)
+
+ fun make_rec_intr T rec_set ((rec_intr_ts, l), (cname, cargs)) =
+ let
+ fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
+ let val free1 = mk_Free "x" U j
+ in (case (strip_dtyp dt, strip_type U) of
+ ((_, DtRec m), (Us, _)) =>
+ let
+ val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k;
+ val i = length Us
+ in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
+ (map (pair "x") Us, List.nth (rec_sets', m) $
+ app_bnds free1 i $ app_bnds free2 i)) :: prems,
+ free1::t1s, free2::t2s)
+ end
+ | _ => (j + 1, k, prems, free1::t1s, t2s))
+ end;
+
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val (_, _, prems, t1s, t2s) = List.foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts)
+
+ in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
+ (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
+ list_comb (List.nth (rec_fns, l), t1s @ t2s)))], l + 1)
+ end;
+
+ val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) =>
+ Library.foldl (make_rec_intr T set_name) (x, #3 (snd d)))
+ (([], 0), descr' ~~ recTs ~~ rec_sets');
+
+ val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+ (map dest_Free rec_fns)
+ (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
+
+ (* prove uniqueness and termination of primrec combinators *)
+
+ val _ = message config "Proving termination and uniqueness of primrec functions ...";
+
+ fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
+ let
+ val distinct_tac =
+ (if i < length newTs then
+ full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1
+ else full_simp_tac dist_ss 1);
+
+ val inject = map (fn r => r RS iffD1)
+ (if i < length newTs then List.nth (constr_inject, i)
+ else #inject (the (Symtab.lookup dt_info tname)));
+
+ fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
+ let
+ val k = length (List.filter is_rec_type cargs)
+
+ in (EVERY [DETERM tac,
+ REPEAT (etac ex1E 1), rtac ex1I 1,
+ DEPTH_SOLVE_1 (ares_tac [intr] 1),
+ REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
+ etac elim 1,
+ REPEAT_DETERM_N j distinct_tac,
+ TRY (dresolve_tac inject 1),
+ REPEAT (etac conjE 1), hyp_subst_tac 1,
+ REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
+ TRY (hyp_subst_tac 1),
+ rtac refl 1,
+ REPEAT_DETERM_N (n - j - 1) distinct_tac],
+ intrs, j + 1)
+ end;
+
+ val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs))
+ ((tac, intrs, 0), constrs);
+
+ in (tac', intrs') end;
+
+ val rec_unique_thms =
+ let
+ val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
+ Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
+ absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
+ (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
+ val cert = cterm_of thy1
+ val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
+ ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
+ val induct' = cterm_instantiate ((map cert induct_Ps) ~~
+ (map cert insts)) induct;
+ val (tac, _) = Library.foldl mk_unique_tac
+ (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
+ THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs),
+ descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
+
+ in split_conj_thm (SkipProof.prove_global thy1 [] []
+ (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
+ end;
+
+ val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
+
+ (* define primrec combinators *)
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.full_bname thy1)
+ (if length descr' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ val (reccomb_defs, thy2) =
+ thy1
+ |> Sign.add_consts_i (map (fn ((name, T), T') =>
+ (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts))
+ |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
+ (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
+ Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+ set $ Free ("x", T) $ Free ("y", T'))))))
+ (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
+ ||> parent_path (#flat_names config)
+ ||> Theory.checkpoint;
+
+
+ (* prove characteristic equations for primrec combinators *)
+
+ val _ = message config "Proving characteristic theorems for primrec combinators ..."
+
+ val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
+ (fn _ => EVERY
+ [rewrite_goals_tac reccomb_defs,
+ rtac the1_equality 1,
+ resolve_tac rec_unique_thms 1,
+ resolve_tac rec_intrs 1,
+ REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
+ (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
+
+ in
+ thy2
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> PureThy.add_thmss [((Binding.name "recs", rec_thms),
+ [Nitpick_Const_Simp_Thms.add])]
+ ||> Sign.parent_path
+ ||> Theory.checkpoint
+ |-> (fn thms => pair (reccomb_names, Library.flat thms))
+ end;
+
+
+(***************************** case combinators *******************************)
+
+fun prove_case_thms (config : config) new_type_names descr sorts reccomb_names primrec_thms thy =
+ let
+ val _ = message config "Proving characteristic theorems for case combinators ...";
+
+ val thy1 = add_path (#flat_names config) (space_implode "_" new_type_names) thy;
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used "'t", HOLogic.typeS);
+
+ fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T';
+
+ val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
+ in Const (@{const_name undefined}, Ts @ Ts' ---> T')
+ end) constrs) descr';
+
+ val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
+
+ (* define case combinators via primrec combinators *)
+
+ val (case_defs, thy2) = Library.foldl (fn ((defs, thy),
+ ((((i, (_, _, constrs)), T), name), recname)) =>
+ let
+ val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
+ val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
+ val frees = Library.take (length cargs, frees');
+ val free = mk_Free "f" (Ts ---> T') j
+ in
+ (free, list_abs_free (map dest_Free frees',
+ list_comb (free, frees)))
+ end) (constrs ~~ (1 upto length constrs)));
+
+ val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
+ val fns = (List.concat (Library.take (i, case_dummy_fns))) @
+ fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
+ val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
+ val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
+ val def = (Binding.name (Long_Name.base_name name ^ "_def"),
+ Logic.mk_equals (list_comb (Const (name, caseT), fns1),
+ list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
+ fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.declare_const [] decl |> snd
+ |> (PureThy.add_defs false o map Thm.no_attributes) [def];
+
+ in (defs @ [def_thm], thy')
+ end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
+ (Library.take (length newTs, reccomb_names)))
+ ||> Theory.checkpoint;
+
+ val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
+ (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
+ (DatatypeProp.make_cases new_type_names descr sorts thy2)
+ in
+ thy2
+ |> Context.the_theory o fold (fold Nitpick_Const_Simp_Thms.add_thm) case_thms
+ o Context.Theory
+ |> parent_path (#flat_names config)
+ |> store_thmss "cases" new_type_names case_thms
+ |-> (fn thmss => pair (thmss, case_names))
+ end;
+
+
+(******************************* case splitting *******************************)
+
+fun prove_split_thms (config : config) new_type_names descr sorts constr_inject dist_rewrites
+ casedist_thms case_thms thy =
+ let
+ val _ = message config "Proving equations for case splitting ...";
+
+ val descr' = flat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
+ exhaustion), case_thms'), T) =
+ let
+ val cert = cterm_of thy;
+ val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
+ val exhaustion' = cterm_instantiate
+ [(cert lhs, cert (Free ("x", T)))] exhaustion;
+ val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
+ (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
+ in
+ (SkipProof.prove_global thy [] [] t1 tacf,
+ SkipProof.prove_global thy [] [] t2 tacf)
+ end;
+
+ val split_thm_pairs = map prove_split_thms
+ ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
+ dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
+
+ val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
+
+ in
+ thy
+ |> store_thms "split" new_type_names split_thms
+ ||>> store_thms "split_asm" new_type_names split_asm_thms
+ |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
+ end;
+
+fun prove_weak_case_congs new_type_names descr sorts thy =
+ let
+ fun prove_weak_case_cong t =
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
+
+ val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
+ new_type_names descr sorts thy)
+
+ in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end;
+
+(************************* additional theorems for TFL ************************)
+
+fun prove_nchotomys (config : config) new_type_names descr sorts casedist_thms thy =
+ let
+ val _ = message config "Proving additional theorems for TFL ...";
+
+ fun prove_nchotomy (t, exhaustion) =
+ let
+ (* For goal i, select the correct disjunct to attack, then prove it *)
+ fun tac i 0 = EVERY [TRY (rtac disjI1 i),
+ hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
+ | tac i n = rtac disjI2 i THEN tac i (n - 1)
+ in
+ SkipProof.prove_global thy [] [] t (fn _ =>
+ EVERY [rtac allI 1,
+ exh_tac (K exhaustion) 1,
+ ALLGOALS (fn i => tac i (i-1))])
+ end;
+
+ val nchotomys =
+ map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
+
+ in thy |> store_thms "nchotomy" new_type_names nchotomys end;
+
+fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
+ let
+ fun prove_case_cong ((t, nchotomy), case_rewrites) =
+ let
+ val (Const ("==>", _) $ tm $ _) = t;
+ val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
+ val cert = cterm_of thy;
+ val nchotomy' = nchotomy RS spec;
+ val [v] = Term.add_vars (concl_of nchotomy') [];
+ val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
+ in
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} =>
+ let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
+ in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
+ cut_facts_tac [nchotomy''] 1,
+ REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
+ REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
+ end)
+ end;
+
+ val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
+ new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
+
+ in thy |> store_thms "case_cong" new_type_names case_congs end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,381 @@
+(* Title: HOL/Tools/datatype_aux.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Auxiliary functions for defining datatypes.
+*)
+
+signature DATATYPE_COMMON =
+sig
+ type config
+ val default_config : config
+ datatype dtyp =
+ DtTFree of string
+ | DtType of string * (dtyp list)
+ | DtRec of int;
+ type descr
+ type info
+end
+
+signature DATATYPE_AUX =
+sig
+ include DATATYPE_COMMON
+
+ val message : config -> string -> unit
+
+ val add_path : bool -> string -> theory -> theory
+ val parent_path : bool -> theory -> theory
+
+ val store_thmss_atts : string -> string list -> attribute list list -> thm list list
+ -> theory -> thm list list * theory
+ val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
+ val store_thms_atts : string -> string list -> attribute list list -> thm list
+ -> theory -> thm list * theory
+ val store_thms : string -> string list -> thm list -> theory -> thm list * theory
+
+ val split_conj_thm : thm -> thm list
+ val mk_conj : term list -> term
+ val mk_disj : term list -> term
+
+ val app_bnds : term -> int -> term
+
+ val cong_tac : int -> tactic
+ val indtac : thm -> string list -> int -> tactic
+ val exh_tac : (string -> thm) -> int -> tactic
+
+ datatype simproc_dist = FewConstrs of thm list
+ | ManyConstrs of thm * simpset;
+
+
+ exception Datatype
+ exception Datatype_Empty of string
+ val name_of_typ : typ -> string
+ val dtyp_of_typ : (string * string list) list -> typ -> dtyp
+ val mk_Free : string -> typ -> int -> term
+ val is_rec_type : dtyp -> bool
+ val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
+ val dest_DtTFree : dtyp -> string
+ val dest_DtRec : dtyp -> int
+ val strip_dtyp : dtyp -> dtyp list * dtyp
+ val body_index : dtyp -> int
+ val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
+ val get_nonrec_types : descr -> (string * sort) list -> typ list
+ val get_branching_types : descr -> (string * sort) list -> typ list
+ val get_arities : descr -> int list
+ val get_rec_types : descr -> (string * sort) list -> typ list
+ val interpret_construction : descr -> (string * sort) list
+ -> { atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a }
+ -> ((string * Term.typ list) * (string * 'a list) list) list
+ val check_nonempty : descr list -> unit
+ val unfold_datatypes :
+ theory -> descr -> (string * sort) list -> info Symtab.table ->
+ descr -> int -> descr list * int
+end;
+
+structure DatatypeAux : DATATYPE_AUX =
+struct
+
+(* datatype option flags *)
+
+type config = { strict: bool, flat_names: bool, quiet: bool };
+val default_config : config =
+ { strict = true, flat_names = false, quiet = false };
+fun message ({ quiet, ...} : config) s =
+ if quiet then () else writeln s;
+
+fun add_path flat_names s = if flat_names then I else Sign.add_path s;
+fun parent_path flat_names = if flat_names then I else Sign.parent_path;
+
+
+(* store theorems in theory *)
+
+fun store_thmss_atts label tnames attss thmss =
+ fold_map (fn ((tname, atts), thms) =>
+ Sign.add_path tname
+ #> PureThy.add_thmss [((Binding.name label, thms), atts)]
+ #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
+ ##> Theory.checkpoint;
+
+fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
+
+fun store_thms_atts label tnames attss thmss =
+ fold_map (fn ((tname, atts), thms) =>
+ Sign.add_path tname
+ #> PureThy.add_thms [((Binding.name label, thms), atts)]
+ #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
+ ##> Theory.checkpoint;
+
+fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
+
+
+(* split theorem thm_1 & ... & thm_n into n theorems *)
+
+fun split_conj_thm th =
+ ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
+
+val mk_conj = foldr1 (HOLogic.mk_binop "op &");
+val mk_disj = foldr1 (HOLogic.mk_binop "op |");
+
+fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
+
+
+fun cong_tac i st = (case Logic.strip_assums_concl
+ (List.nth (prems_of st, i - 1)) of
+ _ $ (_ $ (f $ x) $ (g $ y)) =>
+ let
+ val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
+ val _ $ (_ $ (f' $ x') $ (g' $ y')) =
+ Logic.strip_assums_concl (prop_of cong');
+ val insts = map (pairself (cterm_of (Thm.theory_of_thm st)) o
+ apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
+ apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
+ in compose_tac (false, cterm_instantiate insts cong', 2) i st
+ handle THM _ => no_tac st
+ end
+ | _ => no_tac st);
+
+(* instantiate induction rule *)
+
+fun indtac indrule indnames i st =
+ let
+ val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
+ val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
+ (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
+ val getP = if can HOLogic.dest_imp (hd ts) then
+ (apfst SOME) o HOLogic.dest_imp else pair NONE;
+ val flt = if null indnames then I else
+ filter (fn Free (s, _) => s mem indnames | _ => false);
+ fun abstr (t1, t2) = (case t1 of
+ NONE => (case flt (OldTerm.term_frees t2) of
+ [Free (s, T)] => SOME (absfree (s, T, t2))
+ | _ => NONE)
+ | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
+ val cert = cterm_of (Thm.theory_of_thm st);
+ val insts = List.mapPartial (fn (t, u) => case abstr (getP u) of
+ NONE => NONE
+ | SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u')) (ts ~~ ts');
+ val indrule' = cterm_instantiate insts indrule
+ in
+ rtac indrule' i st
+ end;
+
+(* perform exhaustive case analysis on last parameter of subgoal i *)
+
+fun exh_tac exh_thm_of i state =
+ let
+ val thy = Thm.theory_of_thm state;
+ val prem = nth (prems_of state) (i - 1);
+ val params = Logic.strip_params prem;
+ val (_, Type (tname, _)) = hd (rev params);
+ val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
+ val prem' = hd (prems_of exhaustion);
+ val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
+ val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs),
+ cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t))
+ (Bound 0) params))] exhaustion
+ in compose_tac (false, exhaustion', nprems_of exhaustion) i state
+ end;
+
+(* handling of distinctness theorems *)
+
+datatype simproc_dist = FewConstrs of thm list
+ | ManyConstrs of thm * simpset;
+
+(********************** Internal description of datatypes *********************)
+
+datatype dtyp =
+ DtTFree of string
+ | DtType of string * (dtyp list)
+ | DtRec of int;
+
+(* information about datatypes *)
+
+(* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
+type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
+
+type info =
+ {index : int,
+ alt_names : string list option,
+ descr : descr,
+ sorts : (string * sort) list,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ case_name : string,
+ case_rewrites : thm list,
+ induction : thm,
+ exhaustion : thm,
+ distinct : simproc_dist,
+ inject : thm list,
+ nchotomy : thm,
+ case_cong : thm,
+ weak_case_cong : thm};
+
+fun mk_Free s T i = Free (s ^ (string_of_int i), T);
+
+fun subst_DtTFree _ substs (T as (DtTFree name)) =
+ AList.lookup (op =) substs name |> the_default T
+ | subst_DtTFree i substs (DtType (name, ts)) =
+ DtType (name, map (subst_DtTFree i substs) ts)
+ | subst_DtTFree i _ (DtRec j) = DtRec (i + j);
+
+exception Datatype;
+exception Datatype_Empty of string;
+
+fun dest_DtTFree (DtTFree a) = a
+ | dest_DtTFree _ = raise Datatype;
+
+fun dest_DtRec (DtRec i) = i
+ | dest_DtRec _ = raise Datatype;
+
+fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
+ | is_rec_type (DtRec _) = true
+ | is_rec_type _ = false;
+
+fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
+ | strip_dtyp T = ([], T);
+
+val body_index = dest_DtRec o snd o strip_dtyp;
+
+fun mk_fun_dtyp [] U = U
+ | mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
+
+fun name_of_typ (Type (s, Ts)) =
+ let val s' = Long_Name.base_name s
+ in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
+ [if Syntax.is_identifier s' then s' else "x"])
+ end
+ | name_of_typ _ = "";
+
+fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
+ | dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
+ | dtyp_of_typ new_dts (Type (tname, Ts)) =
+ (case AList.lookup (op =) new_dts tname of
+ NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
+ | SOME vs => if map (try (fst o dest_TFree)) Ts = map SOME vs then
+ DtRec (find_index (curry op = tname o fst) new_dts)
+ else error ("Illegal occurrence of recursive type " ^ tname));
+
+fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
+ | typ_of_dtyp descr sorts (DtRec i) =
+ let val (s, ds, _) = (the o AList.lookup (op =) descr) i
+ in Type (s, map (typ_of_dtyp descr sorts) ds) end
+ | typ_of_dtyp descr sorts (DtType (s, ds)) =
+ Type (s, map (typ_of_dtyp descr sorts) ds);
+
+(* find all non-recursive types in datatype description *)
+
+fun get_nonrec_types descr sorts =
+ map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
+ Library.foldl (fn (Ts', (_, cargs)) =>
+ filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
+
+(* get all recursive types in datatype description *)
+
+fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
+ Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
+
+(* get all branching types *)
+
+fun get_branching_types descr sorts =
+ map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => fold (strip_dtyp #> fst #> fold (insert op =)) cargs)
+ constrs) descr []);
+
+fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
+ (List.filter is_rec_type cargs))) constrs) descr [];
+
+(* interpret construction of datatype *)
+
+fun interpret_construction descr vs { atyp, dtyp } =
+ let
+ val typ_of_dtyp = typ_of_dtyp descr vs;
+ fun interpT dT = case strip_dtyp dT
+ of (dTs, DtRec l) =>
+ let
+ val (tyco, dTs', _) = (the o AList.lookup (op =) descr) l;
+ val Ts = map typ_of_dtyp dTs;
+ val Ts' = map typ_of_dtyp dTs';
+ val is_proper = forall (can dest_TFree) Ts';
+ in dtyp Ts (l, is_proper) (tyco, Ts') end
+ | _ => atyp (typ_of_dtyp dT);
+ fun interpC (c, dTs) = (c, map interpT dTs);
+ fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
+ in map interpD descr end;
+
+(* nonemptiness check for datatypes *)
+
+fun check_nonempty descr =
+ let
+ val descr' = List.concat descr;
+ fun is_nonempty_dt is i =
+ let
+ val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
+ fun arg_nonempty (_, DtRec i) = if i mem is then false
+ else is_nonempty_dt (i::is) i
+ | arg_nonempty _ = true;
+ in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
+ end
+ in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
+ (fn (_, (s, _, _)) => raise Datatype_Empty s)
+ end;
+
+(* unfold a list of mutually recursive datatype specifications *)
+(* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
+(* need to be unfolded *)
+
+fun unfold_datatypes sign orig_descr sorts (dt_info : info Symtab.table) descr i =
+ let
+ fun typ_error T msg = error ("Non-admissible type expression\n" ^
+ Syntax.string_of_typ_global sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
+
+ fun get_dt_descr T i tname dts =
+ (case Symtab.lookup dt_info tname of
+ NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
+ \ nested recursion")
+ | (SOME {index, descr, ...}) =>
+ let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
+ val subst = ((map dest_DtTFree vars) ~~ dts) handle Library.UnequalLengths =>
+ typ_error T ("Type constructor " ^ tname ^ " used with wrong\
+ \ number of arguments")
+ in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
+ (tn, map (subst_DtTFree i subst) args,
+ map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
+ end);
+
+ (* unfold a single constructor argument *)
+
+ fun unfold_arg ((i, Ts, descrs), T) =
+ if is_rec_type T then
+ let val (Us, U) = strip_dtyp T
+ in if exists is_rec_type Us then
+ typ_error T "Non-strictly positive recursive occurrence of type"
+ else (case U of
+ DtType (tname, dts) =>
+ let
+ val (index, descr) = get_dt_descr T i tname dts;
+ val (descr', i') = unfold_datatypes sign orig_descr sorts
+ dt_info descr (i + length descr)
+ in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
+ | _ => (i, Ts @ [T], descrs))
+ end
+ else (i, Ts @ [T], descrs);
+
+ (* unfold a constructor *)
+
+ fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
+ let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
+ in (i', constrs @ [(cname, cargs')], descrs') end;
+
+ (* unfold a single datatype *)
+
+ fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
+ let val (i', constrs', descrs') =
+ Library.foldl unfold_constr ((i, [], descrs), constrs)
+ in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
+ end;
+
+ val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
+
+ in (descr' :: descrs, i') end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_case.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,469 @@
+(* Title: HOL/Tools/datatype_case.ML
+ Author: Konrad Slind, Cambridge University Computer Laboratory
+ Author: Stefan Berghofer, TU Muenchen
+
+Nested case expressions on datatypes.
+*)
+
+signature DATATYPE_CASE =
+sig
+ val make_case: (string -> DatatypeAux.info option) ->
+ Proof.context -> bool -> string list -> term -> (term * term) list ->
+ term * (term * (int * bool)) list
+ val dest_case: (string -> DatatypeAux.info option) -> bool ->
+ string list -> term -> (term * (term * term) list) option
+ val strip_case: (string -> DatatypeAux.info option) -> bool ->
+ term -> (term * (term * term) list) option
+ val case_tr: bool -> (theory -> string -> DatatypeAux.info option)
+ -> Proof.context -> term list -> term
+ val case_tr': (theory -> string -> DatatypeAux.info option) ->
+ string -> Proof.context -> term list -> term
+end;
+
+structure DatatypeCase : DATATYPE_CASE =
+struct
+
+exception CASE_ERROR of string * int;
+
+fun match_type thy pat ob = Sign.typ_match thy (pat, ob) Vartab.empty;
+
+(*---------------------------------------------------------------------------
+ * Get information about datatypes
+ *---------------------------------------------------------------------------*)
+
+fun ty_info (tab : string -> DatatypeAux.info option) s =
+ case tab s of
+ SOME {descr, case_name, index, sorts, ...} =>
+ let
+ val (_, (tname, dts, constrs)) = nth descr index;
+ val mk_ty = DatatypeAux.typ_of_dtyp descr sorts;
+ val T = Type (tname, map mk_ty dts)
+ in
+ SOME {case_name = case_name,
+ constructors = map (fn (cname, dts') =>
+ Const (cname, Logic.varifyT (map mk_ty dts' ---> T))) constrs}
+ end
+ | NONE => NONE;
+
+
+(*---------------------------------------------------------------------------
+ * Each pattern carries with it a tag (i,b) where
+ * i is the clause it came from and
+ * b=true indicates that clause was given by the user
+ * (or is an instantiation of a user supplied pattern)
+ * b=false --> i = ~1
+ *---------------------------------------------------------------------------*)
+
+fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
+
+fun row_of_pat x = fst (snd x);
+
+fun add_row_used ((prfx, pats), (tm, tag)) =
+ fold Term.add_free_names (tm :: pats @ prfx);
+
+(* try to preserve names given by user *)
+fun default_names names ts =
+ map (fn ("", Free (name', _)) => name' | (name, _) => name) (names ~~ ts);
+
+fun strip_constraints (Const ("_constrain", _) $ t $ tT) =
+ strip_constraints t ||> cons tT
+ | strip_constraints t = (t, []);
+
+fun mk_fun_constrain tT t = Syntax.const "_constrain" $ t $
+ (Syntax.free "fun" $ tT $ Syntax.free "dummy");
+
+
+(*---------------------------------------------------------------------------
+ * Produce an instance of a constructor, plus genvars for its arguments.
+ *---------------------------------------------------------------------------*)
+fun fresh_constr ty_match ty_inst colty used c =
+ let
+ val (_, Ty) = dest_Const c
+ val Ts = binder_types Ty;
+ val names = Name.variant_list used
+ (DatatypeProp.make_tnames (map Logic.unvarifyT Ts));
+ val ty = body_type Ty;
+ val ty_theta = ty_match ty colty handle Type.TYPE_MATCH =>
+ raise CASE_ERROR ("type mismatch", ~1)
+ val c' = ty_inst ty_theta c
+ val gvars = map (ty_inst ty_theta o Free) (names ~~ Ts)
+ in (c', gvars)
+ end;
+
+
+(*---------------------------------------------------------------------------
+ * Goes through a list of rows and picks out the ones beginning with a
+ * pattern with constructor = name.
+ *---------------------------------------------------------------------------*)
+fun mk_group (name, T) rows =
+ let val k = length (binder_types T)
+ in fold (fn (row as ((prfx, p :: rst), rhs as (_, (i, _)))) =>
+ fn ((in_group, not_in_group), (names, cnstrts)) => (case strip_comb p of
+ (Const (name', _), args) =>
+ if name = name' then
+ if length args = k then
+ let val (args', cnstrts') = split_list (map strip_constraints args)
+ in
+ ((((prfx, args' @ rst), rhs) :: in_group, not_in_group),
+ (default_names names args', map2 append cnstrts cnstrts'))
+ end
+ else raise CASE_ERROR
+ ("Wrong number of arguments for constructor " ^ name, i)
+ else ((in_group, row :: not_in_group), (names, cnstrts))
+ | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
+ rows (([], []), (replicate k "", replicate k [])) |>> pairself rev
+ end;
+
+(*---------------------------------------------------------------------------
+ * Partition the rows. Not efficient: we should use hashing.
+ *---------------------------------------------------------------------------*)
+fun partition _ _ _ _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
+ | partition ty_match ty_inst type_of used constructors colty res_ty
+ (rows as (((prfx, _ :: rstp), _) :: _)) =
+ let
+ fun part {constrs = [], rows = [], A} = rev A
+ | part {constrs = [], rows = (_, (_, (i, _))) :: _, A} =
+ raise CASE_ERROR ("Not a constructor pattern", i)
+ | part {constrs = c :: crst, rows, A} =
+ let
+ val ((in_group, not_in_group), (names, cnstrts)) =
+ mk_group (dest_Const c) rows;
+ val used' = fold add_row_used in_group used;
+ val (c', gvars) = fresh_constr ty_match ty_inst colty used' c;
+ val in_group' =
+ if null in_group (* Constructor not given *)
+ then
+ let
+ val Ts = map type_of rstp;
+ val xs = Name.variant_list
+ (fold Term.add_free_names gvars used')
+ (replicate (length rstp) "x")
+ in
+ [((prfx, gvars @ map Free (xs ~~ Ts)),
+ (Const ("HOL.undefined", res_ty), (~1, false)))]
+ end
+ else in_group
+ in
+ part{constrs = crst,
+ rows = not_in_group,
+ A = {constructor = c',
+ new_formals = gvars,
+ names = names,
+ constraints = cnstrts,
+ group = in_group'} :: A}
+ end
+ in part {constrs = constructors, rows = rows, A = []}
+ end;
+
+(*---------------------------------------------------------------------------
+ * Misc. routines used in mk_case
+ *---------------------------------------------------------------------------*)
+
+fun mk_pat ((c, c'), l) =
+ let
+ val L = length (binder_types (fastype_of c))
+ fun build (prfx, tag, plist) =
+ let val (args, plist') = chop L plist
+ in (prfx, tag, list_comb (c', args) :: plist') end
+ in map build l end;
+
+fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
+ | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
+
+fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
+ | v_to_pats _ = raise CASE_ERROR ("mk_case: v_to_pats", ~1);
+
+
+(*----------------------------------------------------------------------------
+ * Translation of pattern terms into nested case expressions.
+ *
+ * This performs the translation and also builds the full set of patterns.
+ * Thus it supports the construction of induction theorems even when an
+ * incomplete set of patterns is given.
+ *---------------------------------------------------------------------------*)
+
+fun mk_case tab ctxt ty_match ty_inst type_of used range_ty =
+ let
+ val name = Name.variant used "a";
+ fun expand constructors used ty ((_, []), _) =
+ raise CASE_ERROR ("mk_case: expand_var_row", ~1)
+ | expand constructors used ty (row as ((prfx, p :: rst), rhs)) =
+ if is_Free p then
+ let
+ val used' = add_row_used row used;
+ fun expnd c =
+ let val capp =
+ list_comb (fresh_constr ty_match ty_inst ty used' c)
+ in ((prfx, capp :: rst), pattern_subst [(p, capp)] rhs)
+ end
+ in map expnd constructors end
+ else [row]
+ fun mk {rows = [], ...} = raise CASE_ERROR ("no rows", ~1)
+ | mk {path = [], rows = ((prfx, []), (tm, tag)) :: _} = (* Done *)
+ ([(prfx, tag, [])], tm)
+ | mk {path, rows as ((row as ((_, [Free _]), _)) :: _ :: _)} =
+ mk {path = path, rows = [row]}
+ | mk {path = u :: rstp, rows as ((_, _ :: _), _) :: _} =
+ let val col0 = map (fn ((_, p :: _), (_, (i, _))) => (p, i)) rows
+ in case Option.map (apfst head_of)
+ (find_first (not o is_Free o fst) col0) of
+ NONE =>
+ let
+ val rows' = map (fn ((v, _), row) => row ||>
+ pattern_subst [(v, u)] |>> v_to_prfx) (col0 ~~ rows);
+ val (pref_patl, tm) = mk {path = rstp, rows = rows'}
+ in (map v_to_pats pref_patl, tm) end
+ | SOME (Const (cname, cT), i) => (case ty_info tab cname of
+ NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ cname, i)
+ | SOME {case_name, constructors} =>
+ let
+ val pty = body_type cT;
+ val used' = fold Term.add_free_names rstp used;
+ val nrows = maps (expand constructors used' pty) rows;
+ val subproblems = partition ty_match ty_inst type_of used'
+ constructors pty range_ty nrows;
+ val new_formals = map #new_formals subproblems
+ val constructors' = map #constructor subproblems
+ val news = map (fn {new_formals, group, ...} =>
+ {path = new_formals @ rstp, rows = group}) subproblems;
+ val (pat_rect, dtrees) = split_list (map mk news);
+ val case_functions = map2
+ (fn {new_formals, names, constraints, ...} =>
+ fold_rev (fn ((x as Free (_, T), s), cnstrts) => fn t =>
+ Abs (if s = "" then name else s, T,
+ abstract_over (x, t)) |>
+ fold mk_fun_constrain cnstrts)
+ (new_formals ~~ names ~~ constraints))
+ subproblems dtrees;
+ val types = map type_of (case_functions @ [u]);
+ val case_const = Const (case_name, types ---> range_ty)
+ val tree = list_comb (case_const, case_functions @ [u])
+ val pat_rect1 = flat (map mk_pat
+ (constructors ~~ constructors' ~~ pat_rect))
+ in (pat_rect1, tree)
+ end)
+ | SOME (t, i) => raise CASE_ERROR ("Not a datatype constructor: " ^
+ Syntax.string_of_term ctxt t, i)
+ end
+ | mk _ = raise CASE_ERROR ("Malformed row matrix", ~1)
+ in mk
+ end;
+
+fun case_error s = error ("Error in case expression:\n" ^ s);
+
+(* Repeated variable occurrences in a pattern are not allowed. *)
+fun no_repeat_vars ctxt pat = fold_aterms
+ (fn x as Free (s, _) => (fn xs =>
+ if member op aconv xs x then
+ case_error (quote s ^ " occurs repeatedly in the pattern " ^
+ quote (Syntax.string_of_term ctxt pat))
+ else x :: xs)
+ | _ => I) pat [];
+
+fun gen_make_case ty_match ty_inst type_of tab ctxt err used x clauses =
+ let
+ fun string_of_clause (pat, rhs) = Syntax.string_of_term ctxt
+ (Syntax.const "_case1" $ pat $ rhs);
+ val _ = map (no_repeat_vars ctxt o fst) clauses;
+ val rows = map_index (fn (i, (pat, rhs)) =>
+ (([], [pat]), (rhs, (i, true)))) clauses;
+ val rangeT = (case distinct op = (map (type_of o snd) clauses) of
+ [] => case_error "no clauses given"
+ | [T] => T
+ | _ => case_error "all cases must have the same result type");
+ val used' = fold add_row_used rows used;
+ val (patts, case_tm) = mk_case tab ctxt ty_match ty_inst type_of
+ used' rangeT {path = [x], rows = rows}
+ handle CASE_ERROR (msg, i) => case_error (msg ^
+ (if i < 0 then ""
+ else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
+ val patts1 = map
+ (fn (_, tag, [pat]) => (pat, tag)
+ | _ => case_error "error in pattern-match translation") patts;
+ val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1
+ val finals = map row_of_pat patts2
+ val originals = map (row_of_pat o #2) rows
+ val _ = case originals \\ finals of
+ [] => ()
+ | is => (if err then case_error else warning)
+ ("The following clauses are redundant (covered by preceding clauses):\n" ^
+ cat_lines (map (string_of_clause o nth clauses) is));
+ in
+ (case_tm, patts2)
+ end;
+
+fun make_case tab ctxt = gen_make_case
+ (match_type (ProofContext.theory_of ctxt)) Envir.subst_TVars fastype_of tab ctxt;
+val make_case_untyped = gen_make_case (K (K Vartab.empty))
+ (K (Term.map_types (K dummyT))) (K dummyT);
+
+
+(* parse translation *)
+
+fun case_tr err tab_of ctxt [t, u] =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ (* replace occurrences of dummy_pattern by distinct variables *)
+ (* internalize constant names *)
+ fun prep_pat ((c as Const ("_constrain", _)) $ t $ tT) used =
+ let val (t', used') = prep_pat t used
+ in (c $ t' $ tT, used') end
+ | prep_pat (Const ("dummy_pattern", T)) used =
+ let val x = Name.variant used "x"
+ in (Free (x, T), x :: used) end
+ | prep_pat (Const (s, T)) used =
+ (case try (unprefix Syntax.constN) s of
+ SOME c => (Const (c, T), used)
+ | NONE => (Const (Sign.intern_const thy s, T), used))
+ | prep_pat (v as Free (s, T)) used =
+ let val s' = Sign.intern_const thy s
+ in
+ if Sign.declared_const thy s' then
+ (Const (s', T), used)
+ else (v, used)
+ end
+ | prep_pat (t $ u) used =
+ let
+ val (t', used') = prep_pat t used;
+ val (u', used'') = prep_pat u used'
+ in
+ (t' $ u', used'')
+ end
+ | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
+ fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
+ let val (l', cnstrts) = strip_constraints l
+ in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
+ end
+ | dest_case1 t = case_error "dest_case1";
+ fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
+ | dest_case2 t = [t];
+ val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
+ val (case_tm, _) = make_case_untyped (tab_of thy) ctxt err []
+ (fold (fn tT => fn t => Syntax.const "_constrain" $ t $ tT)
+ (flat cnstrts) t) cases;
+ in case_tm end
+ | case_tr _ _ _ ts = case_error "case_tr";
+
+
+(*---------------------------------------------------------------------------
+ * Pretty printing of nested case expressions
+ *---------------------------------------------------------------------------*)
+
+(* destruct one level of pattern matching *)
+
+fun gen_dest_case name_of type_of tab d used t =
+ case apfst name_of (strip_comb t) of
+ (SOME cname, ts as _ :: _) =>
+ let
+ val (fs, x) = split_last ts;
+ fun strip_abs i t =
+ let
+ val zs = strip_abs_vars t;
+ val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
+ val (xs, ys) = chop i zs;
+ val u = list_abs (ys, strip_abs_body t);
+ val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
+ (map fst xs) ~~ map snd xs)
+ in (xs', subst_bounds (rev xs', u)) end;
+ fun is_dependent i t =
+ let val k = length (strip_abs_vars t) - i
+ in k < 0 orelse exists (fn j => j >= k)
+ (loose_bnos (strip_abs_body t))
+ end;
+ fun count_cases (_, _, true) = I
+ | count_cases (c, (_, body), false) =
+ AList.map_default op aconv (body, []) (cons c);
+ val is_undefined = name_of #> equal (SOME "HOL.undefined");
+ fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
+ in case ty_info tab cname of
+ SOME {constructors, case_name} =>
+ if length fs = length constructors then
+ let
+ val cases = map (fn (Const (s, U), t) =>
+ let
+ val k = length (binder_types U);
+ val p as (xs, _) = strip_abs k t
+ in
+ (Const (s, map type_of xs ---> type_of x),
+ p, is_dependent k t)
+ end) (constructors ~~ fs);
+ val cases' = sort (int_ord o swap o pairself (length o snd))
+ (fold_rev count_cases cases []);
+ val R = type_of t;
+ val dummy = if d then Const ("dummy_pattern", R)
+ else Free (Name.variant used "x", R)
+ in
+ SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
+ SOME (_, cs) =>
+ if length cs = length constructors then [hd cases]
+ else filter_out (fn (_, (_, body), _) => is_undefined body) cases
+ | NONE => case cases' of
+ [] => cases
+ | (default, cs) :: _ =>
+ if length cs = 1 then cases
+ else if length cs = length constructors then
+ [hd cases, (dummy, ([], default), false)]
+ else
+ filter_out (fn (c, _, _) => member op aconv cs c) cases @
+ [(dummy, ([], default), false)]))
+ end handle CASE_ERROR _ => NONE
+ else NONE
+ | _ => NONE
+ end
+ | _ => NONE;
+
+val dest_case = gen_dest_case (try (dest_Const #> fst)) fastype_of;
+val dest_case' = gen_dest_case
+ (try (dest_Const #> fst #> unprefix Syntax.constN)) (K dummyT);
+
+
+(* destruct nested patterns *)
+
+fun strip_case'' dest (pat, rhs) =
+ case dest (Term.add_free_names pat []) rhs of
+ SOME (exp as Free _, clauses) =>
+ if member op aconv (OldTerm.term_frees pat) exp andalso
+ not (exists (fn (_, rhs') =>
+ member op aconv (OldTerm.term_frees rhs') exp) clauses)
+ then
+ maps (strip_case'' dest) (map (fn (pat', rhs') =>
+ (subst_free [(exp, pat')] pat, rhs')) clauses)
+ else [(pat, rhs)]
+ | _ => [(pat, rhs)];
+
+fun gen_strip_case dest t = case dest [] t of
+ SOME (x, clauses) =>
+ SOME (x, maps (strip_case'' dest) clauses)
+ | NONE => NONE;
+
+val strip_case = gen_strip_case oo dest_case;
+val strip_case' = gen_strip_case oo dest_case';
+
+
+(* print translation *)
+
+fun case_tr' tab_of cname ctxt ts =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val consts = ProofContext.consts_of ctxt;
+ fun mk_clause (pat, rhs) =
+ let val xs = Term.add_frees pat []
+ in
+ Syntax.const "_case1" $
+ map_aterms
+ (fn Free p => Syntax.mark_boundT p
+ | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
+ | t => t) pat $
+ map_aterms
+ (fn x as Free (s, T) =>
+ if member (op =) xs (s, T) then Syntax.mark_bound s else x
+ | t => t) rhs
+ end
+ in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
+ SOME (x, clauses) => Syntax.const "_case_syntax" $ x $
+ foldr1 (fn (t, u) => Syntax.const "_case2" $ t $ u)
+ (map mk_clause clauses)
+ | NONE => raise Match
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,455 @@
+(* Title: HOL/Tools/datatype_codegen.ML
+ Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
+
+Code generator facilities for inductive datatypes.
+*)
+
+signature DATATYPE_CODEGEN =
+sig
+ val find_shortest_path: Datatype.descr -> int -> (string * int) option
+ val mk_eq_eqns: theory -> string -> (thm * bool) list
+ val mk_case_cert: theory -> string -> thm
+ val setup: theory -> theory
+end;
+
+structure DatatypeCodegen : DATATYPE_CODEGEN =
+struct
+
+(** find shortest path to constructor with no recursive arguments **)
+
+fun find_nonempty (descr: Datatype.descr) is i =
+ let
+ val (_, _, constrs) = the (AList.lookup (op =) descr i);
+ fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i
+ then NONE
+ else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i)
+ | arg_nonempty _ = SOME 0;
+ fun max xs = Library.foldl
+ (fn (NONE, _) => NONE
+ | (SOME i, SOME j) => SOME (Int.max (i, j))
+ | (_, NONE) => NONE) (SOME 0, xs);
+ val xs = sort (int_ord o pairself snd)
+ (map_filter (fn (s, dts) => Option.map (pair s)
+ (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
+ in case xs of [] => NONE | x :: _ => SOME x end;
+
+fun find_shortest_path descr i = find_nonempty descr [i] i;
+
+
+(** SML code generator **)
+
+open Codegen;
+
+(* datatype definition *)
+
+fun add_dt_defs thy defs dep module (descr: Datatype.descr) sorts gr =
+ let
+ val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
+ val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
+ exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
+
+ val (_, (tname, _, _)) :: _ = descr';
+ val node_id = tname ^ " (type)";
+ val module' = if_library (thyname_of_type thy tname) module;
+
+ fun mk_dtdef prfx [] gr = ([], gr)
+ | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr =
+ let
+ val tvs = map DatatypeAux.dest_DtTFree dts;
+ val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
+ val ((_, type_id), gr') = mk_type_id module' tname gr;
+ val (ps, gr'') = gr' |>
+ fold_map (fn (cname, cargs) =>
+ fold_map (invoke_tycodegen thy defs node_id module' false)
+ cargs ##>>
+ mk_const_id module' cname) cs';
+ val (rest, gr''') = mk_dtdef "and " xs gr''
+ in
+ (Pretty.block (str prfx ::
+ (if null tvs then [] else
+ [mk_tuple (map str tvs), str " "]) @
+ [str (type_id ^ " ="), Pretty.brk 1] @
+ List.concat (separate [Pretty.brk 1, str "| "]
+ (map (fn (ps', (_, cname)) => [Pretty.block
+ (str cname ::
+ (if null ps' then [] else
+ List.concat ([str " of", Pretty.brk 1] ::
+ separate [str " *", Pretty.brk 1]
+ (map single ps'))))]) ps))) :: rest, gr''')
+ end;
+
+ fun mk_constr_term cname Ts T ps =
+ List.concat (separate [str " $", Pretty.brk 1]
+ ([str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
+ mk_type false (Ts ---> T), str ")"] :: ps));
+
+ fun mk_term_of_def gr prfx [] = []
+ | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) =
+ let
+ val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
+ val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val T = Type (tname, dts');
+ val rest = mk_term_of_def gr "and " xs;
+ val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx =>
+ let val args = map (fn i =>
+ str ("x" ^ string_of_int i)) (1 upto length Ts)
+ in (Pretty.blk (4,
+ [str prfx, mk_term_of gr module' false T, Pretty.brk 1,
+ if null Ts then str (snd (get_const_id gr cname))
+ else parens (Pretty.block
+ [str (snd (get_const_id gr cname)),
+ Pretty.brk 1, mk_tuple args]),
+ str " =", Pretty.brk 1] @
+ mk_constr_term cname Ts T
+ (map2 (fn x => fn U => [Pretty.block [mk_term_of gr module' false U,
+ Pretty.brk 1, x]]) args Ts)), " | ")
+ end) cs' prfx
+ in eqs @ rest end;
+
+ fun mk_gen_of_def gr prfx [] = []
+ | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
+ let
+ val tvs = map DatatypeAux.dest_DtTFree dts;
+ val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val T = Type (tname, Us);
+ val (cs1, cs2) =
+ List.partition (exists DatatypeAux.is_rec_type o snd) cs;
+ val SOME (cname, _) = find_shortest_path descr i;
+
+ fun mk_delay p = Pretty.block
+ [str "fn () =>", Pretty.brk 1, p];
+
+ fun mk_force p = Pretty.block [p, Pretty.brk 1, str "()"];
+
+ fun mk_constr s b (cname, dts) =
+ let
+ val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s
+ (DatatypeAux.typ_of_dtyp descr sorts dt))
+ [str (if b andalso DatatypeAux.is_rec_type dt then "0"
+ else "j")]) dts;
+ val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val xs = map str
+ (DatatypeProp.indexify_names (replicate (length dts) "x"));
+ val ts = map str
+ (DatatypeProp.indexify_names (replicate (length dts) "t"));
+ val (_, id) = get_const_id gr cname
+ in
+ mk_let
+ (map2 (fn p => fn q => mk_tuple [p, q]) xs ts ~~ gs)
+ (mk_tuple
+ [case xs of
+ _ :: _ :: _ => Pretty.block
+ [str id, Pretty.brk 1, mk_tuple xs]
+ | _ => mk_app false (str id) xs,
+ mk_delay (Pretty.block (mk_constr_term cname Ts T
+ (map (single o mk_force) ts)))])
+ end;
+
+ fun mk_choice [c] = mk_constr "(i-1)" false c
+ | mk_choice cs = Pretty.block [str "one_of",
+ Pretty.brk 1, Pretty.blk (1, str "[" ::
+ List.concat (separate [str ",", Pretty.fbrk]
+ (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
+ [str "]"]), Pretty.brk 1, str "()"];
+
+ val gs = maps (fn s =>
+ let val s' = strip_tname s
+ in [str (s' ^ "G"), str (s' ^ "T")] end) tvs;
+ val gen_name = "gen_" ^ snd (get_type_id gr tname)
+
+ in
+ Pretty.blk (4, separate (Pretty.brk 1)
+ (str (prfx ^ gen_name ^
+ (if null cs1 then "" else "'")) :: gs @
+ (if null cs1 then [] else [str "i"]) @
+ [str "j"]) @
+ [str " =", Pretty.brk 1] @
+ (if not (null cs1) andalso not (null cs2)
+ then [str "frequency", Pretty.brk 1,
+ Pretty.blk (1, [str "[",
+ mk_tuple [str "i", mk_delay (mk_choice cs1)],
+ str ",", Pretty.fbrk,
+ mk_tuple [str "1", mk_delay (mk_choice cs2)],
+ str "]"]), Pretty.brk 1, str "()"]
+ else if null cs2 then
+ [Pretty.block [str "(case", Pretty.brk 1,
+ str "i", Pretty.brk 1, str "of",
+ Pretty.brk 1, str "0 =>", Pretty.brk 1,
+ mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
+ Pretty.brk 1, str "| _ =>", Pretty.brk 1,
+ mk_choice cs1, str ")"]]
+ else [mk_choice cs2])) ::
+ (if null cs1 then []
+ else [Pretty.blk (4, separate (Pretty.brk 1)
+ (str ("and " ^ gen_name) :: gs @ [str "i"]) @
+ [str " =", Pretty.brk 1] @
+ separate (Pretty.brk 1) (str (gen_name ^ "'") :: gs @
+ [str "i", str "i"]))]) @
+ mk_gen_of_def gr "and " xs
+ end
+
+ in
+ (module', (add_edge_acyclic (node_id, dep) gr
+ handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
+ let
+ val gr1 = add_edge (node_id, dep)
+ (new_node (node_id, (NONE, "", "")) gr);
+ val (dtdef, gr2) = mk_dtdef "datatype " descr' gr1 ;
+ in
+ map_node node_id (K (NONE, module',
+ string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
+ [str ";"])) ^ "\n\n" ^
+ (if "term_of" mem !mode then
+ string_of (Pretty.blk (0, separate Pretty.fbrk
+ (mk_term_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
+ else "") ^
+ (if "test" mem !mode then
+ string_of (Pretty.blk (0, separate Pretty.fbrk
+ (mk_gen_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
+ else ""))) gr2
+ end)
+ end;
+
+
+(* case expressions *)
+
+fun pretty_case thy defs dep module brack constrs (c as Const (_, T)) ts gr =
+ let val i = length constrs
+ in if length ts <= i then
+ invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr
+ else
+ let
+ val ts1 = Library.take (i, ts);
+ val t :: ts2 = Library.drop (i, ts);
+ val names = List.foldr OldTerm.add_term_names
+ (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1;
+ val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
+
+ fun pcase [] [] [] gr = ([], gr)
+ | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr =
+ let
+ val j = length cargs;
+ val xs = Name.variant_list names (replicate j "x");
+ val Us' = Library.take (j, fst (strip_type U));
+ val frees = map Free (xs ~~ Us');
+ val (cp, gr0) = invoke_codegen thy defs dep module false
+ (list_comb (Const (cname, Us' ---> dT), frees)) gr;
+ val t' = Envir.beta_norm (list_comb (t, frees));
+ val (p, gr1) = invoke_codegen thy defs dep module false t' gr0;
+ val (ps, gr2) = pcase cs ts Us gr1;
+ in
+ ([Pretty.block [cp, str " =>", Pretty.brk 1, p]] :: ps, gr2)
+ end;
+
+ val (ps1, gr1) = pcase constrs ts1 Ts gr ;
+ val ps = List.concat (separate [Pretty.brk 1, str "| "] ps1);
+ val (p, gr2) = invoke_codegen thy defs dep module false t gr1;
+ val (ps2, gr3) = fold_map (invoke_codegen thy defs dep module true) ts2 gr2;
+ in ((if not (null ts2) andalso brack then parens else I)
+ (Pretty.block (separate (Pretty.brk 1)
+ (Pretty.block ([str "(case ", p, str " of",
+ Pretty.brk 1] @ ps @ [str ")"]) :: ps2))), gr3)
+ end
+ end;
+
+
+(* constructors *)
+
+fun pretty_constr thy defs dep module brack args (c as Const (s, T)) ts gr =
+ let val i = length args
+ in if i > 1 andalso length ts < i then
+ invoke_codegen thy defs dep module brack (eta_expand c ts i) gr
+ else
+ let
+ val id = mk_qual_id module (get_const_id gr s);
+ val (ps, gr') = fold_map
+ (invoke_codegen thy defs dep module (i = 1)) ts gr;
+ in (case args of
+ _ :: _ :: _ => (if brack then parens else I)
+ (Pretty.block [str id, Pretty.brk 1, mk_tuple ps])
+ | _ => (mk_app brack (str id) ps), gr')
+ end
+ end;
+
+
+(* code generators for terms and types *)
+
+fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
+ (c as Const (s, T), ts) =>
+ (case Datatype.datatype_of_case thy s of
+ SOME {index, descr, ...} =>
+ if is_some (get_assoc_code thy (s, T)) then NONE else
+ SOME (pretty_case thy defs dep module brack
+ (#3 (the (AList.lookup op = descr index))) c ts gr )
+ | NONE => case (Datatype.datatype_of_constr thy s, strip_type T) of
+ (SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
+ if is_some (get_assoc_code thy (s, T)) then NONE else
+ let
+ val SOME (tyname', _, constrs) = AList.lookup op = descr index;
+ val SOME args = AList.lookup op = constrs s
+ in
+ if tyname <> tyname' then NONE
+ else SOME (pretty_constr thy defs
+ dep module brack args c ts (snd (invoke_tycodegen thy defs dep module false U gr)))
+ end
+ | _ => NONE)
+ | _ => NONE);
+
+fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
+ (case Datatype.get_datatype thy s of
+ NONE => NONE
+ | SOME {descr, sorts, ...} =>
+ if is_some (get_assoc_type thy s) then NONE else
+ let
+ val (ps, gr') = fold_map
+ (invoke_tycodegen thy defs dep module false) Ts gr;
+ val (module', gr'') = add_dt_defs thy defs dep module descr sorts gr' ;
+ val (tyid, gr''') = mk_type_id module' s gr''
+ in SOME (Pretty.block ((if null Ts then [] else
+ [mk_tuple ps, str " "]) @
+ [str (mk_qual_id module tyid)]), gr''')
+ end)
+ | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
+
+
+(** generic code generator **)
+
+(* liberal addition of code data for datatypes *)
+
+fun mk_constr_consts thy vs dtco cos =
+ let
+ val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+ val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
+ in if is_some (try (Code.constrset_of_consts thy) cs')
+ then SOME cs
+ else NONE
+ end;
+
+
+(* case certificates *)
+
+fun mk_case_cert thy tyco =
+ let
+ val raw_thms =
+ (#case_rewrites o Datatype.the_datatype thy) tyco;
+ val thms as hd_thm :: _ = raw_thms
+ |> Conjunction.intr_balanced
+ |> Thm.unvarify
+ |> Conjunction.elim_balanced (length raw_thms)
+ |> map Simpdata.mk_meta_eq
+ |> map Drule.zero_var_indexes
+ val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
+ | _ => I) (Thm.prop_of hd_thm) [];
+ val rhs = hd_thm
+ |> Thm.prop_of
+ |> Logic.dest_equals
+ |> fst
+ |> Term.strip_comb
+ |> apsnd (fst o split_last)
+ |> list_comb;
+ val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
+ val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
+ in
+ thms
+ |> Conjunction.intr_balanced
+ |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
+ |> Thm.implies_intr asm
+ |> Thm.generalize ([], params) 0
+ |> AxClass.unoverload thy
+ |> Thm.varifyT
+ end;
+
+
+(* equality *)
+
+fun mk_eq_eqns thy dtco =
+ let
+ val (vs, cos) = Datatype.the_datatype_spec thy dtco;
+ val { descr, index, inject = inject_thms, ... } = Datatype.the_datatype thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ $ t1 $ t2;
+ fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
+ fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
+ val triv_injects = map_filter
+ (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
+ | _ => NONE) cos;
+ fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
+ trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
+ val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
+ fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
+ [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
+ val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
+ val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
+ val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
+ addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
+ addsimprocs [Datatype.distinct_simproc]);
+ fun prove prop = SkipProof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
+ |> Simpdata.mk_eq;
+ in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
+
+fun add_equality vs dtcos thy =
+ let
+ fun add_def dtco lthy =
+ let
+ val ty = Type (dtco, map TFree vs);
+ fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
+ $ Free ("x", ty) $ Free ("y", ty);
+ val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
+ val def' = Syntax.check_term lthy def;
+ val ((_, (_, thm)), lthy') = Specification.definition
+ (NONE, (Attrib.empty_binding, def')) lthy;
+ val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
+ val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
+ in (thm', lthy') end;
+ fun tac thms = Class.intro_classes_tac []
+ THEN ALLGOALS (ProofContext.fact_tac thms);
+ fun add_eq_thms dtco thy =
+ let
+ val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
+ val thy_ref = Theory.check_thy thy;
+ fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
+ in
+ Code.add_eqnl (const, Lazy.lazy mk_thms) thy
+ end;
+ in
+ thy
+ |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
+ |> fold_map add_def dtcos
+ |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
+ (fn _ => fn def_thms => tac def_thms) def_thms)
+ |-> (fn def_thms => fold Code.del_eqn def_thms)
+ |> fold add_eq_thms dtcos
+ end;
+
+
+(* register a datatype etc. *)
+
+fun add_all_code config dtcos thy =
+ let
+ val (vs :: _, coss) = (split_list o map (Datatype.the_datatype_spec thy)) dtcos;
+ val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+ val css = if exists is_none any_css then []
+ else map_filter I any_css;
+ val case_rewrites = maps (#case_rewrites o Datatype.the_datatype thy) dtcos;
+ val certs = map (mk_case_cert thy) dtcos;
+ in
+ if null css then thy
+ else thy
+ |> tap (fn _ => DatatypeAux.message config "Registering datatype for code generator ...")
+ |> fold Code.add_datatype css
+ |> fold_rev Code.add_default_eqn case_rewrites
+ |> fold Code.add_case certs
+ |> add_equality vs dtcos
+ end;
+
+
+(** theory setup **)
+
+val setup =
+ add_codegen "datatype" datatype_codegen
+ #> add_tycodegen "datatype" datatype_tycodegen
+ #> Datatype.interpretation add_all_code
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,435 @@
+(* Title: HOL/Tools/datatype_prop.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Characteristic properties of datatypes.
+*)
+
+signature DATATYPE_PROP =
+sig
+ val indexify_names: string list -> string list
+ val make_tnames: typ list -> string list
+ val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list
+ val make_distincts : DatatypeAux.descr list ->
+ (string * sort) list -> (int * term list) list (*no symmetric inequalities*)
+ val make_ind : DatatypeAux.descr list -> (string * sort) list -> term
+ val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list
+ val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list ->
+ string list -> typ list * typ list
+ val make_primrecs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_cases : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list list
+ val make_splits : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> (term * term) list
+ val make_weak_case_congs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_case_congs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_nchotomys : DatatypeAux.descr list ->
+ (string * sort) list -> term list
+end;
+
+structure DatatypeProp : DATATYPE_PROP =
+struct
+
+open DatatypeAux;
+
+fun indexify_names names =
+ let
+ fun index (x :: xs) tab =
+ (case AList.lookup (op =) tab x of
+ NONE => if member (op =) xs x then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
+ | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
+ | index [] _ = [];
+ in index names [] end;
+
+fun make_tnames Ts =
+ let
+ fun type_name (TFree (name, _)) = implode (tl (explode name))
+ | type_name (Type (name, _)) =
+ let val name' = Long_Name.base_name name
+ in if Syntax.is_identifier name' then name' else "x" end;
+ in indexify_names (map type_name Ts) end;
+
+
+(************************* injectivity of constructors ************************)
+
+fun make_injs descr sorts =
+ let
+ val descr' = flat descr;
+ fun make_inj T (cname, cargs) =
+ if null cargs then I else
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val constr_t = Const (cname, Ts ---> T);
+ val tnames = make_tnames Ts;
+ val frees = map Free (tnames ~~ Ts);
+ val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
+ in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
+ foldr1 (HOLogic.mk_binop "op &")
+ (map HOLogic.mk_eq (frees ~~ frees')))))
+ end;
+ in
+ map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
+ (hd descr) (Library.take (length (hd descr), get_rec_types descr' sorts))
+ end;
+
+
+(************************* distinctness of constructors ***********************)
+
+fun make_distincts descr sorts =
+ let
+ val descr' = flat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun prep_constr (cname, cargs) = (cname, map (typ_of_dtyp descr' sorts) cargs);
+
+ fun make_distincts' _ [] = []
+ | make_distincts' T ((cname, cargs)::constrs) =
+ let
+ val frees = map Free ((make_tnames cargs) ~~ cargs);
+ val t = list_comb (Const (cname, cargs ---> T), frees);
+
+ fun make_distincts'' (cname', cargs') =
+ let
+ val frees' = map Free ((map ((op ^) o (rpair "'"))
+ (make_tnames cargs')) ~~ cargs');
+ val t' = list_comb (Const (cname', cargs' ---> T), frees')
+ in
+ HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.mk_eq (t, t'))
+ end
+
+ in map make_distincts'' constrs @ make_distincts' T constrs end;
+
+ in
+ map2 (fn ((_, (_, _, constrs))) => fn T =>
+ (length constrs, make_distincts' T (map prep_constr constrs))) (hd descr) newTs
+ end;
+
+
+(********************************* induction **********************************)
+
+fun make_ind descr sorts =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val pnames = if length descr' = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr');
+
+ fun make_pred i T =
+ let val T' = T --> HOLogic.boolT
+ in Free (List.nth (pnames, i), T') end;
+
+ fun make_ind_prem k T (cname, cargs) =
+ let
+ fun mk_prem ((dt, s), T) =
+ let val (Us, U) = strip_type T
+ in list_all (map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us)))
+ end;
+
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr' sorts) recs;
+ val tnames = Name.variant_list pnames (make_tnames Ts);
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = tnames ~~ Ts;
+ val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
+
+ in list_all_free (frees, Logic.list_implies (prems,
+ HOLogic.mk_Trueprop (make_pred k T $
+ list_comb (Const (cname, Ts ---> T), map Free frees))))
+ end;
+
+ val prems = List.concat (map (fn ((i, (_, _, constrs)), T) =>
+ map (make_ind_prem i T) constrs) (descr' ~~ recTs));
+ val tnames = make_tnames recTs;
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
+ (descr' ~~ recTs ~~ tnames)))
+
+ in Logic.list_implies (prems, concl) end;
+
+(******************************* case distinction *****************************)
+
+fun make_casedists descr sorts =
+ let
+ val descr' = List.concat descr;
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees
+ in list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
+ HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))))
+ end;
+
+ fun make_casedist ((_, (_, _, constrs)), T) =
+ let val prems = map (make_casedist_prem T) constrs
+ in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)))
+ end
+
+ in map make_casedist
+ ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts))
+ end;
+
+(*************** characteristic equations for primrec combinator **************)
+
+fun make_primrec_Ts descr sorts used =
+ let
+ val descr' = List.concat descr;
+
+ val rec_result_Ts = map TFree (Name.variant_list used (replicate (length descr') "'t") ~~
+ replicate (length descr') HOLogic.typeS);
+
+ val reccomb_fn_Ts = List.concat (map (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts);
+
+ fun mk_argT (dt, T) =
+ binder_types T ---> List.nth (rec_result_Ts, body_index dt);
+
+ val argTs = Ts @ map mk_argT recs
+ in argTs ---> List.nth (rec_result_Ts, i)
+ end) constrs) descr');
+
+ in (rec_result_Ts, reccomb_fn_Ts) end;
+
+fun make_primrecs new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+
+ val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.intern_const thy)
+ (if length descr' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) =
+ let
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr' sorts) recs;
+ val tnames = make_tnames Ts;
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = map Free (tnames ~~ Ts);
+ val frees' = map Free (rec_tnames ~~ recTs');
+
+ fun mk_reccomb ((dt, T), t) =
+ let val (Us, U) = strip_type T
+ in list_abs (map (pair "x") Us,
+ List.nth (reccombs, body_index dt) $ app_bnds t (length Us))
+ end;
+
+ val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees')
+
+ in (ts @ [HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
+ list_comb (f, frees @ reccombs')))], fs)
+ end
+
+ in fst (Library.foldl (fn (x, ((dt, T), comb_t)) =>
+ Library.foldl (make_primrec T comb_t) (x, #3 (snd dt)))
+ (([], rec_fns), descr' ~~ recTs ~~ reccombs))
+ end;
+
+(****************** make terms of form t_case f1 ... fn *********************)
+
+fun make_case_combs new_type_names descr sorts thy fname =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used "'t", HOLogic.typeS);
+
+ val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let val Ts = map (typ_of_dtyp descr' sorts) cargs
+ in Ts ---> T' end) constrs) (hd descr);
+
+ val case_names = map (fn s =>
+ Sign.intern_const thy (s ^ "_case")) new_type_names
+ in
+ map (fn ((name, Ts), T) => list_comb
+ (Const (name, Ts @ [T] ---> T'),
+ map (uncurry (mk_Free fname)) (Ts ~~ (1 upto length Ts))))
+ (case_names ~~ case_fn_Ts ~~ newTs)
+ end;
+
+(**************** characteristic equations for case combinator ****************)
+
+fun make_cases new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun make_case T comb_t ((cname, cargs), f) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = map Free ((make_tnames Ts) ~~ Ts)
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
+ list_comb (f, frees)))
+ end
+
+ in map (fn (((_, (_, _, constrs)), T), comb_t) =>
+ map (make_case T comb_t) (constrs ~~ (snd (strip_comb comb_t))))
+ ((hd descr) ~~ newTs ~~ (make_case_combs new_type_names descr sorts thy "f"))
+ end;
+
+
+(*************************** the "split" - equations **************************)
+
+fun make_splits new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used' = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
+ val P = Free ("P", T' --> HOLogic.boolT);
+
+ fun make_split (((_, (_, _, constrs)), T), comb_t) =
+ let
+ val (_, fs) = strip_comb comb_t;
+ val used = ["P", "x"] @ (map (fst o dest_Free) fs);
+
+ fun process_constr (((cname, cargs), f), (t1s, t2s)) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts);
+ val eqn = HOLogic.mk_eq (Free ("x", T),
+ list_comb (Const (cname, Ts ---> T), frees));
+ val P' = P $ list_comb (f, frees)
+ in ((List.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t))
+ (HOLogic.imp $ eqn $ P') frees)::t1s,
+ (List.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t))
+ (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) frees)::t2s)
+ end;
+
+ val (t1s, t2s) = List.foldr process_constr ([], []) (constrs ~~ fs);
+ val lhs = P $ (comb_t $ Free ("x", T))
+ in
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ mk_disj t2s)))
+ end
+
+ in map make_split ((hd descr) ~~ newTs ~~
+ (make_case_combs new_type_names descr sorts thy "f"))
+ end;
+
+(************************* additional rules for TFL ***************************)
+
+fun make_weak_case_congs new_type_names descr sorts thy =
+ let
+ val case_combs = make_case_combs new_type_names descr sorts thy "f";
+
+ fun mk_case_cong comb =
+ let
+ val Type ("fun", [T, _]) = fastype_of comb;
+ val M = Free ("M", T);
+ val M' = Free ("M'", T);
+ in
+ Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb $ M')))
+ end
+ in
+ map mk_case_cong case_combs
+ end;
+
+
+(*---------------------------------------------------------------------------
+ * Structure of case congruence theorem looks like this:
+ *
+ * (M = M')
+ * ==> (!!x1,...,xk. (M' = C1 x1..xk) ==> (f1 x1..xk = g1 x1..xk))
+ * ==> ...
+ * ==> (!!x1,...,xj. (M' = Cn x1..xj) ==> (fn x1..xj = gn x1..xj))
+ * ==>
+ * (ty_case f1..fn M = ty_case g1..gn M')
+ *---------------------------------------------------------------------------*)
+
+fun make_case_congs new_type_names descr sorts thy =
+ let
+ val case_combs = make_case_combs new_type_names descr sorts thy "f";
+ val case_combs' = make_case_combs new_type_names descr sorts thy "g";
+
+ fun mk_case_cong ((comb, comb'), (_, (_, _, constrs))) =
+ let
+ val Type ("fun", [T, _]) = fastype_of comb;
+ val (_, fs) = strip_comb comb;
+ val (_, gs) = strip_comb comb';
+ val used = ["M", "M'"] @ map (fst o dest_Free) (fs @ gs);
+ val M = Free ("M", T);
+ val M' = Free ("M'", T);
+
+ fun mk_clause ((f, g), (cname, _)) =
+ let
+ val (Ts, _) = strip_type (fastype_of f);
+ val tnames = Name.variant_list used (make_tnames Ts);
+ val frees = map Free (tnames ~~ Ts)
+ in
+ list_all_free (tnames ~~ Ts, Logic.mk_implies
+ (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (M', list_comb (Const (cname, Ts ---> T), frees))),
+ HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (list_comb (f, frees), list_comb (g, frees)))))
+ end
+
+ in
+ Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')) ::
+ map mk_clause (fs ~~ gs ~~ constrs),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb' $ M')))
+ end
+
+ in
+ map mk_case_cong (case_combs ~~ case_combs' ~~ hd descr)
+ end;
+
+(*---------------------------------------------------------------------------
+ * Structure of exhaustion theorem looks like this:
+ *
+ * !v. (? y1..yi. v = C1 y1..yi) | ... | (? y1..yj. v = Cn y1..yj)
+ *---------------------------------------------------------------------------*)
+
+fun make_nchotomys descr sorts =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun mk_eqn T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val tnames = Name.variant_list ["v"] (make_tnames Ts);
+ val frees = tnames ~~ Ts
+ in
+ List.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t))
+ (HOLogic.mk_eq (Free ("v", T),
+ list_comb (Const (cname, Ts ---> T), map Free frees))) frees
+ end
+
+ in map (fn ((_, (_, _, constrs)), T) =>
+ HOLogic.mk_Trueprop (HOLogic.mk_all ("v", T, mk_disj (map (mk_eqn T) constrs))))
+ (hd descr ~~ newTs)
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,230 @@
+(* Title: HOL/Tools/datatype_realizer.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Porgram extraction from proofs involving datatypes:
+Realizers for induction and case analysis
+*)
+
+signature DATATYPE_REALIZER =
+sig
+ val add_dt_realizers: Datatype.config -> string list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure DatatypeRealizer : DATATYPE_REALIZER =
+struct
+
+open DatatypeAux;
+
+fun subsets i j = if i <= j then
+ let val is = subsets (i+1) j
+ in map (fn ks => i::ks) is @ is end
+ else [[]];
+
+fun forall_intr_prf (t, prf) =
+ let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
+ in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
+
+fun prf_of thm =
+ Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
+
+fun prf_subst_vars inst =
+ Proofterm.map_proof_terms (subst_vars ([], inst)) I;
+
+fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
+
+fun tname_of (Type (s, _)) = s
+ | tname_of _ = "";
+
+fun mk_realizes T = Const ("realizes", T --> HOLogic.boolT --> HOLogic.boolT);
+
+fun make_ind sorts ({descr, rec_names, rec_rewrites, induction, ...} : info) is thy =
+ let
+ val recTs = get_rec_types descr sorts;
+ val pnames = if length descr = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
+
+ val rec_result_Ts = map (fn ((i, _), P) =>
+ if i mem is then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
+ (descr ~~ pnames);
+
+ fun make_pred i T U r x =
+ if i mem is then
+ Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
+ else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
+
+ fun mk_all i s T t =
+ if i mem is then list_all_free ([(s, T)], t) else t;
+
+ val (prems, rec_fns) = split_list (flat (fst (fold_map
+ (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
+ val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
+ val frees = tnames ~~ Ts;
+
+ fun mk_prems vs [] =
+ let
+ val rT = nth (rec_result_Ts) i;
+ val vs' = filter_out is_unit vs;
+ val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
+ val f' = Envir.eta_contract (list_abs_free
+ (map dest_Free vs, if i mem is then list_comb (f, vs')
+ else HOLogic.unit));
+ in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
+ (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
+ end
+ | mk_prems vs (((dt, s), T) :: ds) =
+ let
+ val k = body_index dt;
+ val (Us, U) = strip_type T;
+ val i = length Us;
+ val rT = nth (rec_result_Ts) k;
+ val r = Free ("r" ^ s, Us ---> rT);
+ val (p, f) = mk_prems (vs @ [r]) ds
+ in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
+ (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred k rT U (app_bnds r i)
+ (app_bnds (Free (s, T)) i))), p)), f)
+ end
+
+ in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
+ constrs) (descr ~~ recTs) 1)));
+
+ fun mk_proj j [] t = t
+ | mk_proj j (i :: is) t = if null is then t else
+ if (j: int) = i then HOLogic.mk_fst t
+ else mk_proj j is (HOLogic.mk_snd t);
+
+ val tnames = DatatypeProp.make_tnames recTs;
+ val fTs = map fastype_of rec_fns;
+ val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
+ (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
+ (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
+ val r = if null is then Extraction.nullt else
+ foldr1 HOLogic.mk_prod (List.mapPartial (fn (((((i, _), T), U), s), tname) =>
+ if i mem is then SOME
+ (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
+ else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), U), tname) =>
+ make_pred i U T (mk_proj i is r) (Free (tname, T)))
+ (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
+ val cert = cterm_of thy;
+ val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
+ (HOLogic.dest_Trueprop (concl_of induction))) ~~ ps);
+
+ val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
+ (fn prems =>
+ [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
+ rtac (cterm_instantiate inst induction) 1,
+ ALLGOALS ObjectLogic.atomize_prems_tac,
+ rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
+ REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
+ REPEAT (etac allE i) THEN atac i)) 1)]);
+
+ val ind_name = Thm.get_name induction;
+ val vs = map (fn i => List.nth (pnames, i)) is;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> PureThy.store_thm
+ (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
+ ||> Sign.restore_naming thy;
+
+ val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []);
+ val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
+ val ivs1 = map Var (filter_out (fn (_, T) =>
+ tname_of (body_type T) mem ["set", "bool"]) ivs);
+ val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (AList.lookup (op =) rvs ixn))) ivs;
+
+ val prf = List.foldr forall_intr_prf
+ (List.foldr (fn ((f, p), prf) =>
+ (case head_of (strip_abs_body f) of
+ Free (s, T) =>
+ let val T' = Logic.varifyT T
+ in Abst (s, SOME T', Proofterm.prf_abstract_over
+ (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
+ end
+ | _ => AbsP ("H", SOME p, prf)))
+ (Proofterm.proof_combP
+ (prf_of thm', map PBound (length prems - 1 downto 0))) (rec_fns ~~ prems_of thm)) ivs2;
+
+ val r' = if null is then r else Logic.varify (List.foldr (uncurry lambda)
+ r (map Logic.unvarify ivs1 @ filter_out is_unit
+ (map (head_of o strip_abs_body) rec_fns)));
+
+ in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
+
+
+fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaustion, ...} : info) thy =
+ let
+ val cert = cterm_of thy;
+ val rT = TFree ("'P", HOLogic.typeS);
+ val rT' = TVar (("'P", 0), HOLogic.typeS);
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees;
+ val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
+ in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
+ HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
+ list_comb (r, free_ts)))))
+ end;
+
+ val SOME (_, _, constrs) = AList.lookup (op =) descr index;
+ val T = List.nth (get_rec_types descr sorts, index);
+ val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
+ val r = Const (case_name, map fastype_of rs ---> T --> rT);
+
+ val y = Var (("y", 0), Logic.legacy_varifyT T);
+ val y' = Free ("y", T);
+
+ val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
+ HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
+ list_comb (r, rs @ [y'])))))
+ (fn prems =>
+ [rtac (cterm_instantiate [(cert y, cert y')] exhaustion) 1,
+ ALLGOALS (EVERY'
+ [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
+ resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
+
+ val exh_name = Thm.get_name exhaustion;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
+ ||> Sign.restore_naming thy;
+
+ val P = Var (("P", 0), rT' --> HOLogic.boolT);
+ val prf = forall_intr_prf (y, forall_intr_prf (P,
+ List.foldr (fn ((p, r), prf) =>
+ forall_intr_prf (Logic.legacy_varify r, AbsP ("H", SOME (Logic.varify p),
+ prf))) (Proofterm.proof_combP (prf_of thm',
+ map PBound (length prems - 1 downto 0))) (prems ~~ rs)));
+ val r' = Logic.legacy_varify (Abs ("y", Logic.legacy_varifyT T,
+ list_abs (map dest_Free rs, list_comb (r,
+ map Bound ((length rs - 1 downto 0) @ [length rs])))));
+
+ in Extraction.add_realizers_i
+ [(exh_name, (["P"], r', prf)),
+ (exh_name, ([], Extraction.nullt, prf_of exhaustion))] thy'
+ end;
+
+fun add_dt_realizers config names thy =
+ if ! Proofterm.proofs < 2 then thy
+ else let
+ val _ = message config "Adding realizers for induction and case analysis ..."
+ val infos = map (Datatype.the_datatype thy) names;
+ val info :: _ = infos;
+ in
+ thy
+ |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
+ |> fold_rev (make_casedists (#sorts info)) infos
+ end;
+
+val setup = Datatype.interpretation add_dt_realizers;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,643 @@
+(* Title: HOL/Tools/datatype_rep_proofs.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Definitional introduction of datatypes
+Proof of characteristic theorems:
+
+ - injectivity of constructors
+ - distinctness of constructors
+ - induction theorem
+*)
+
+signature DATATYPE_REP_PROOFS =
+sig
+ include DATATYPE_COMMON
+ val distinctness_limit : int Config.T
+ val distinctness_limit_setup : theory -> theory
+ val representation_proofs : config -> info Symtab.table ->
+ string list -> descr list -> (string * sort) list ->
+ (binding * mixfix) list -> (binding * mixfix) list list -> attribute
+ -> theory -> (thm list list * thm list list * thm list list *
+ DatatypeAux.simproc_dist list * thm) * theory
+end;
+
+structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
+struct
+
+open DatatypeAux;
+
+(*the kind of distinctiveness axioms depends on number of constructors*)
+val (distinctness_limit, distinctness_limit_setup) =
+ Attrib.config_int "datatype_distinctness_limit" 7;
+
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+
+val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
+
+
+(** theory context references **)
+
+val f_myinv_f = thm "f_myinv_f";
+val myinv_f_f = thm "myinv_f_f";
+
+
+fun exh_thm_of (dt_info : info Symtab.table) tname =
+ #exhaustion (the (Symtab.lookup dt_info tname));
+
+(******************************************************************************)
+
+fun representation_proofs (config : config) (dt_info : info Symtab.table)
+ new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
+ let
+ val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
+ val node_name = "Datatype.node";
+ val In0_name = "Datatype.In0";
+ val In1_name = "Datatype.In1";
+ val Scons_name = "Datatype.Scons";
+ val Leaf_name = "Datatype.Leaf";
+ val Numb_name = "Datatype.Numb";
+ val Lim_name = "Datatype.Lim";
+ val Suml_name = "Datatype.Suml";
+ val Sumr_name = "Datatype.Sumr";
+
+ val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
+ In0_eq, In1_eq, In0_not_In1, In1_not_In0,
+ Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy)
+ ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
+ "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
+ "Lim_inject", "Suml_inject", "Sumr_inject"];
+
+ val descr' = flat descr;
+
+ val big_name = space_implode "_" new_type_names;
+ val thy1 = add_path (#flat_names config) big_name thy;
+ val big_rec_name = big_name ^ "_rep_set";
+ val rep_set_names' =
+ (if length descr' = 1 then [big_rec_name] else
+ (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
+
+ val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
+ val leafTs' = get_nonrec_types descr' sorts;
+ val branchTs = get_branching_types descr' sorts;
+ val branchT = if null branchTs then HOLogic.unitT
+ else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
+ val arities = get_arities descr' \ 0;
+ val unneeded_vars = hd tyvars \\ List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
+ val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+ val oldTs = Library.drop (length (hd descr), recTs);
+ val sumT = if null leafTs then HOLogic.unitT
+ else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
+ val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
+ val UnivT = HOLogic.mk_setT Univ_elT;
+ val UnivT' = Univ_elT --> HOLogic.boolT;
+ val Collect = Const ("Collect", UnivT' --> UnivT);
+
+ val In0 = Const (In0_name, Univ_elT --> Univ_elT);
+ val In1 = Const (In1_name, Univ_elT --> Univ_elT);
+ val Leaf = Const (Leaf_name, sumT --> Univ_elT);
+ val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
+
+ (* make injections needed for embedding types in leaves *)
+
+ fun mk_inj T' x =
+ let
+ fun mk_inj' T n i =
+ if n = 1 then x else
+ let val n2 = n div 2;
+ val Type (_, [T1, T2]) = T
+ in
+ if i <= n2 then
+ Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
+ else
+ Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
+ end
+ in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
+ end;
+
+ (* make injections for constructors *)
+
+ fun mk_univ_inj ts = BalancedTree.access
+ {left = fn t => In0 $ t,
+ right = fn t => In1 $ t,
+ init =
+ if ts = [] then Const (@{const_name undefined}, Univ_elT)
+ else foldr1 (HOLogic.mk_binop Scons_name) ts};
+
+ (* function spaces *)
+
+ fun mk_fun_inj T' x =
+ let
+ fun mk_inj T n i =
+ if n = 1 then x else
+ let
+ val n2 = n div 2;
+ val Type (_, [T1, T2]) = T;
+ fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
+ in
+ if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
+ else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
+ end
+ in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
+ end;
+
+ val mk_lim = List.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t)));
+
+ (************** generate introduction rules for representing set **********)
+
+ val _ = message config "Constructing representing sets ...";
+
+ (* make introduction rule for a single constructor *)
+
+ fun make_intr s n (i, (_, cargs)) =
+ let
+ fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of
+ (dts, DtRec k) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) dts;
+ val free_t =
+ app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
+ in (j + 1, list_all (map (pair "x") Ts,
+ HOLogic.mk_Trueprop
+ (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
+ mk_lim free_t Ts :: ts)
+ end
+ | _ =>
+ let val T = typ_of_dtyp descr' sorts dt
+ in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
+ end);
+
+ val (_, prems, ts) = List.foldr mk_prem (1, [], []) cargs;
+ val concl = HOLogic.mk_Trueprop
+ (Free (s, UnivT') $ mk_univ_inj ts n i)
+ in Logic.list_implies (prems, concl)
+ end;
+
+ val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
+ map (make_intr rep_set_name (length constrs))
+ ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
+
+ val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
+ (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
+
+ (********************************* typedef ********************************)
+
+ val (typedefs, thy3) = thy2 |>
+ parent_path (#flat_names config) |>
+ fold_map (fn ((((name, mx), tvs), c), name') =>
+ Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
+ (Collect $ Const (c, UnivT')) NONE
+ (rtac exI 1 THEN rtac CollectI 1 THEN
+ QUIET_BREADTH_FIRST (has_fewer_prems 1)
+ (resolve_tac rep_intrs 1)))
+ (types_syntax ~~ tyvars ~~
+ (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
+ add_path (#flat_names config) big_name;
+
+ (*********************** definition of constructors ***********************)
+
+ val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
+ val rep_names = map (curry op ^ "Rep_") new_type_names;
+ val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
+ (1 upto (length (flat (tl descr))));
+ val all_rep_names = map (Sign.intern_const thy3) rep_names @
+ map (Sign.full_bname thy3) rep_names';
+
+ (* isomorphism declarations *)
+
+ val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
+ (oldTs ~~ rep_names');
+
+ (* constructor definitions *)
+
+ fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
+ let
+ fun constr_arg (dt, (j, l_args, r_args)) =
+ let val T = typ_of_dtyp descr' sorts dt;
+ val free_t = mk_Free "x" T j
+ in (case (strip_dtyp dt, strip_type T) of
+ ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
+ (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
+ app_bnds free_t (length Us)) Us :: r_args)
+ | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
+ end;
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
+ val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
+ val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
+ val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
+ val lhs = list_comb (Const (cname, constrT), l_args);
+ val rhs = mk_univ_inj r_args n i;
+ val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
+ val def_name = Long_Name.base_name cname ^ "_def";
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.add_consts_i [(cname', constrT, mx)]
+ |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
+
+ in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
+
+ (* constructor definitions for datatype *)
+
+ fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
+ ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
+ let
+ val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
+ val rep_const = cterm_of thy
+ (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
+ val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
+ val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+ val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
+ ((add_path (#flat_names config) tname thy, defs, [], 1), constrs ~~ constr_syntax)
+ in
+ (parent_path (#flat_names config) thy', defs', eqns @ [eqns'],
+ rep_congs @ [cong'], dist_lemmas @ [dist])
+ end;
+
+ val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
+ ((thy3 |> Sign.add_consts_i iso_decls |> parent_path (#flat_names config), [], [], [], []),
+ hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
+
+ (*********** isomorphisms for new types (introduced by typedef) ***********)
+
+ val _ = message config "Proving isomorphism properties ...";
+
+ val newT_iso_axms = map (fn (_, td) =>
+ (collect_simp (#Abs_inverse td), #Rep_inverse td,
+ collect_simp (#Rep td))) typedefs;
+
+ val newT_iso_inj_thms = map (fn (_, td) =>
+ (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
+
+ (********* isomorphisms between existing types and "unfolded" types *******)
+
+ (*---------------------------------------------------------------------*)
+ (* isomorphisms are defined using primrec-combinators: *)
+ (* generate appropriate functions for instantiating primrec-combinator *)
+ (* *)
+ (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
+ (* *)
+ (* also generate characteristic equations for isomorphisms *)
+ (* *)
+ (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
+ (*---------------------------------------------------------------------*)
+
+ fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
+ let
+ val argTs = map (typ_of_dtyp descr' sorts) cargs;
+ val T = List.nth (recTs, k);
+ val rep_name = List.nth (all_rep_names, k);
+ val rep_const = Const (rep_name, T --> Univ_elT);
+ val constr = Const (cname, argTs ---> T);
+
+ fun process_arg ks' ((i2, i2', ts, Ts), dt) =
+ let
+ val T' = typ_of_dtyp descr' sorts dt;
+ val (Us, U) = strip_type T'
+ in (case strip_dtyp dt of
+ (_, DtRec j) => if j mem ks' then
+ (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
+ (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
+ Ts @ [Us ---> Univ_elT])
+ else
+ (i2 + 1, i2', ts @ [mk_lim
+ (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
+ app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
+ | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
+ end;
+
+ val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
+ val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
+ val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
+ val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
+
+ val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
+
+ in (fs @ [f], eqns @ [eqn], i + 1) end;
+
+ (* define isomorphisms for all mutually recursive datatypes in list ds *)
+
+ fun make_iso_defs (ds, (thy, char_thms)) =
+ let
+ val ks = map fst ds;
+ val (_, (tname, _, _)) = hd ds;
+ val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
+
+ fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
+ let
+ val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
+ ((fs, eqns, 1), constrs);
+ val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
+ in (fs', eqns', isos @ [iso]) end;
+
+ val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
+ val fTs = map fastype_of fs;
+ val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
+ Logic.mk_equals (Const (iso_name, T --> Univ_elT),
+ list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
+ val (def_thms, thy') =
+ apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy);
+
+ (* prove characteristic equations *)
+
+ val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
+ val char_thms' = map (fn eqn => SkipProof.prove_global thy' [] [] eqn
+ (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
+
+ in (thy', char_thms' @ char_thms) end;
+
+ val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs
+ (add_path (#flat_names config) big_name thy4, []) (tl descr));
+
+ (* prove isomorphism properties *)
+
+ fun mk_funs_inv thy thm =
+ let
+ val prop = Thm.prop_of thm;
+ val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
+ (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
+ val used = OldTerm.add_term_tfree_names (a, []);
+
+ fun mk_thm i =
+ let
+ val Ts = map (TFree o rpair HOLogic.typeS)
+ (Name.variant_list used (replicate i "'t"));
+ val f = Free ("f", Ts ---> U)
+ in SkipProof.prove_global thy [] [] (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.list_all
+ (map (pair "x") Ts, S $ app_bnds f i)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
+ r $ (a $ app_bnds f i)), f))))
+ (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
+ REPEAT (etac allE 1), rtac thm 1, atac 1])
+ end
+ in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
+
+ (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *)
+
+ val fun_congs = map (fn T => make_elim (Drule.instantiate'
+ [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
+
+ fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
+ let
+ val (_, (tname, _, _)) = hd ds;
+ val {induction, ...} = the (Symtab.lookup dt_info tname);
+
+ fun mk_ind_concl (i, _) =
+ let
+ val T = List.nth (recTs, i);
+ val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
+ val rep_set_name = List.nth (rep_set_names, i)
+ in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
+ HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
+ HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
+ Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
+ end;
+
+ val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
+
+ val rewrites = map mk_meta_eq iso_char_thms;
+ val inj_thms' = map snd newT_iso_inj_thms @
+ map (fn r => r RS @{thm injD}) inj_thms;
+
+ val inj_thm = SkipProof.prove_global thy5 [] []
+ (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
+ [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ REPEAT (EVERY
+ [rtac allI 1, rtac impI 1,
+ exh_tac (exh_thm_of dt_info) 1,
+ REPEAT (EVERY
+ [hyp_subst_tac 1,
+ rewrite_goals_tac rewrites,
+ REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
+ (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
+ ORELSE (EVERY
+ [REPEAT (eresolve_tac (Scons_inject ::
+ map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
+ REPEAT (cong_tac 1), rtac refl 1,
+ REPEAT (atac 1 ORELSE (EVERY
+ [REPEAT (rtac ext 1),
+ REPEAT (eresolve_tac (mp :: allE ::
+ map make_elim (Suml_inject :: Sumr_inject ::
+ Lim_inject :: inj_thms') @ fun_congs) 1),
+ atac 1]))])])])]);
+
+ val inj_thms'' = map (fn r => r RS @{thm datatype_injI})
+ (split_conj_thm inj_thm);
+
+ val elem_thm =
+ SkipProof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
+ (fn _ =>
+ EVERY [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ rewrite_goals_tac rewrites,
+ REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
+ ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
+
+ in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
+ end;
+
+ val (iso_inj_thms_unfolded, iso_elem_thms) = List.foldr prove_iso_thms
+ ([], map #3 newT_iso_axms) (tl descr);
+ val iso_inj_thms = map snd newT_iso_inj_thms @
+ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
+
+ (* prove dt_rep_set_i x --> x : range dt_Rep_i *)
+
+ fun mk_iso_t (((set_name, iso_name), i), T) =
+ let val isoT = T --> Univ_elT
+ in HOLogic.imp $
+ (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
+ (if i < length newTs then HOLogic.true_const
+ else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
+ Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
+ Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
+ end;
+
+ val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
+ (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
+
+ (* all the theorems are proved by one single simultaneous induction *)
+
+ val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq}))
+ iso_inj_thms_unfolded;
+
+ val iso_thms = if length descr = 1 then [] else
+ Library.drop (length newTs, split_conj_thm
+ (SkipProof.prove_global thy5 [] [] iso_t (fn _ => EVERY
+ [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ REPEAT (rtac TrueI 1),
+ rewrite_goals_tac (mk_meta_eq choice_eq ::
+ symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
+ rewrite_goals_tac (map symmetric range_eqs),
+ REPEAT (EVERY
+ [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
+ maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
+ TRY (hyp_subst_tac 1),
+ rtac (sym RS range_eqI) 1,
+ resolve_tac iso_char_thms 1])])));
+
+ val Abs_inverse_thms' =
+ map #1 newT_iso_axms @
+ map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
+ iso_inj_thms_unfolded iso_thms;
+
+ val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
+
+ (******************* freeness theorems for constructors *******************)
+
+ val _ = message config "Proving freeness of constructors ...";
+
+ (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *)
+
+ fun prove_constr_rep_thm eqn =
+ let
+ val inj_thms = map fst newT_iso_inj_thms;
+ val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
+ in SkipProof.prove_global thy5 [] [] eqn (fn _ => EVERY
+ [resolve_tac inj_thms 1,
+ rewrite_goals_tac rewrites,
+ rtac refl 3,
+ resolve_tac rep_intrs 2,
+ REPEAT (resolve_tac iso_elem_thms 1)])
+ end;
+
+ (*--------------------------------------------------------------*)
+ (* constr_rep_thms and rep_congs are used to prove distinctness *)
+ (* of constructors. *)
+ (*--------------------------------------------------------------*)
+
+ val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
+
+ val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
+ dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
+ (constr_rep_thms ~~ dist_lemmas);
+
+ fun prove_distinct_thms _ _ (_, []) = []
+ | prove_distinct_thms lim dist_rewrites' (k, ts as _ :: _) =
+ if k >= lim then [] else let
+ (*number of constructors < distinctness_limit : C_i ... ~= C_j ...*)
+ fun prove [] = []
+ | prove (t :: ts) =
+ let
+ val dist_thm = SkipProof.prove_global thy5 [] [] t (fn _ =>
+ EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
+ in dist_thm :: standard (dist_thm RS not_sym) :: prove ts end;
+ in prove ts end;
+
+ val distinct_thms = DatatypeProp.make_distincts descr sorts
+ |> map2 (prove_distinct_thms
+ (Config.get_thy thy5 distinctness_limit)) dist_rewrites;
+
+ val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
+ if length constrs < Config.get_thy thy5 distinctness_limit
+ then FewConstrs dists
+ else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
+ constr_rep_thms ~~ rep_congs ~~ distinct_thms);
+
+ (* prove injectivity of constructors *)
+
+ fun prove_constr_inj_thm rep_thms t =
+ let val inj_thms = Scons_inject :: (map make_elim
+ (iso_inj_thms @
+ [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
+ Lim_inject, Suml_inject, Sumr_inject]))
+ in SkipProof.prove_global thy5 [] [] t (fn _ => EVERY
+ [rtac iffI 1,
+ REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
+ dresolve_tac rep_congs 1, dtac box_equals 1,
+ REPEAT (resolve_tac rep_thms 1),
+ REPEAT (eresolve_tac inj_thms 1),
+ REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
+ REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
+ atac 1]))])
+ end;
+
+ val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
+ ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
+
+ val ((constr_inject', distinct_thms'), thy6) =
+ thy5
+ |> parent_path (#flat_names config)
+ |> store_thmss "inject" new_type_names constr_inject
+ ||>> store_thmss "distinct" new_type_names distinct_thms;
+
+ (*************************** induction theorem ****************************)
+
+ val _ = message config "Proving induction rule for datatypes ...";
+
+ val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
+ (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
+ val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
+
+ fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
+ let
+ val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
+ mk_Free "x" T i;
+
+ val Abs_t = if i < length newTs then
+ Const (Sign.intern_const thy6
+ ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
+ else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
+ Const (List.nth (all_rep_names, i), T --> Univ_elT)
+
+ in (prems @ [HOLogic.imp $
+ (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
+ (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
+ concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
+ end;
+
+ val (indrule_lemma_prems, indrule_lemma_concls) =
+ Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
+
+ val cert = cterm_of thy6;
+
+ val indrule_lemma = SkipProof.prove_global thy6 [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
+ [REPEAT (etac conjE 1),
+ REPEAT (EVERY
+ [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
+ etac mp 1, resolve_tac iso_elem_thms 1])]);
+
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
+ map (Free o apfst fst o dest_Var) Ps;
+ val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
+
+ val dt_induct_prop = DatatypeProp.make_ind descr sorts;
+ val dt_induct = SkipProof.prove_global thy6 []
+ (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
+ (fn {prems, ...} => EVERY
+ [rtac indrule_lemma' 1,
+ (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ EVERY (map (fn (prem, r) => (EVERY
+ [REPEAT (eresolve_tac Abs_inverse_thms 1),
+ simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
+ DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
+ (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
+
+ val ([dt_induct'], thy7) =
+ thy6
+ |> Sign.add_path big_name
+ |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
+ ||> Sign.parent_path
+ ||> Theory.checkpoint;
+
+ in
+ ((constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct'), thy7)
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/auto_term.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,36 @@
+(* Title: HOL/Tools/Function/auto_term.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Method "relation" to commence a termination proof using a user-specified relation.
+*)
+
+signature FUNDEF_RELATION =
+sig
+ val relation_tac: Proof.context -> term -> int -> tactic
+ val setup: theory -> theory
+end
+
+structure FundefRelation : FUNDEF_RELATION =
+struct
+
+fun inst_thm ctxt rel st =
+ let
+ val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+ val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
+ val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
+ val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
+ in
+ Drule.cterm_instantiate [(Rvar, rel')] st'
+ end
+
+fun relation_tac ctxt rel i =
+ TRY (FundefCommon.apply_termination_rule ctxt i)
+ THEN PRIMITIVE (inst_thm ctxt rel)
+
+val setup =
+ Method.setup @{binding relation}
+ (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
+ "proves termination using a user-specified wellfounded relation"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/context_tree.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,278 @@
+(* Title: HOL/Tools/Function/context_tree.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Builds and traverses trees of nested contexts along a term.
+*)
+
+signature FUNDEF_CTXTREE =
+sig
+ type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
+ type ctx_tree
+
+ (* FIXME: This interface is a mess and needs to be cleaned up! *)
+ val get_fundef_congs : Proof.context -> thm list
+ val add_fundef_cong : thm -> Context.generic -> Context.generic
+ val map_fundef_congs : (thm list -> thm list) -> Context.generic -> Context.generic
+
+ val cong_add: attribute
+ val cong_del: attribute
+
+ val mk_tree: (string * typ) -> term -> Proof.context -> term -> ctx_tree
+
+ val inst_tree: theory -> term -> term -> ctx_tree -> ctx_tree
+
+ val export_term : ctxt -> term -> term
+ val export_thm : theory -> ctxt -> thm -> thm
+ val import_thm : theory -> ctxt -> thm -> thm
+
+ val traverse_tree :
+ (ctxt -> term ->
+ (ctxt * thm) list ->
+ (ctxt * thm) list * 'b ->
+ (ctxt * thm) list * 'b)
+ -> ctx_tree -> 'b -> 'b
+
+ val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list
+end
+
+structure FundefCtxTree : FUNDEF_CTXTREE =
+struct
+
+type ctxt = (string * typ) list * thm list
+
+open FundefCommon
+open FundefLib
+
+structure FundefCongs = GenericDataFun
+(
+ type T = thm list
+ val empty = []
+ val extend = I
+ fun merge _ = Thm.merge_thms
+);
+
+val get_fundef_congs = FundefCongs.get o Context.Proof
+val map_fundef_congs = FundefCongs.map
+val add_fundef_cong = FundefCongs.map o Thm.add_thm
+
+(* congruence rules *)
+
+val cong_add = Thm.declaration_attribute (map_fundef_congs o Thm.add_thm o safe_mk_meta_eq);
+val cong_del = Thm.declaration_attribute (map_fundef_congs o Thm.del_thm o safe_mk_meta_eq);
+
+
+type depgraph = int IntGraph.T
+
+datatype ctx_tree
+ = Leaf of term
+ | Cong of (thm * depgraph * (ctxt * ctx_tree) list)
+ | RCall of (term * ctx_tree)
+
+
+(* Maps "Trueprop A = B" to "A" *)
+val rhs_of = snd o HOLogic.dest_eq o HOLogic.dest_Trueprop
+
+
+(*** Dependency analysis for congruence rules ***)
+
+fun branch_vars t =
+ let
+ val t' = snd (dest_all_all t)
+ val (assumes, concl) = Logic.strip_horn t'
+ in (fold Term.add_vars assumes [], Term.add_vars concl [])
+ end
+
+fun cong_deps crule =
+ let
+ val num_branches = map_index (apsnd branch_vars) (prems_of crule)
+ in
+ IntGraph.empty
+ |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches
+ |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) =>
+ if i = j orelse null (c1 inter t2)
+ then I else IntGraph.add_edge_acyclic (i,j))
+ num_branches num_branches
+ end
+
+val default_congs = map (fn c => c RS eq_reflection) [@{thm "cong"}, @{thm "ext"}]
+
+
+
+(* Called on the INSTANTIATED branches of the congruence rule *)
+fun mk_branch ctx t =
+ let
+ val (ctx', fixes, impl) = dest_all_all_ctx ctx t
+ val (assms, concl) = Logic.strip_horn impl
+ in
+ (ctx', fixes, assms, rhs_of concl)
+ end
+
+fun find_cong_rule ctx fvar h ((r,dep)::rs) t =
+ (let
+ val thy = ProofContext.theory_of ctx
+
+ val tt' = Logic.mk_equals (Pattern.rewrite_term thy [(Free fvar, h)] [] t, t)
+ val (c, subs) = (concl_of r, prems_of r)
+
+ val subst = Pattern.match (ProofContext.theory_of ctx) (c, tt') (Vartab.empty, Vartab.empty)
+ val branches = map (mk_branch ctx o Envir.beta_norm o Envir.subst_vars subst) subs
+ val inst = map (fn v => (cterm_of thy (Var v), cterm_of thy (Envir.subst_vars subst (Var v)))) (Term.add_vars c [])
+ in
+ (cterm_instantiate inst r, dep, branches)
+ end
+ handle Pattern.MATCH => find_cong_rule ctx fvar h rs t)
+ | find_cong_rule _ _ _ [] _ = sys_error "Function/context_tree.ML: No cong rule found!"
+
+
+fun mk_tree fvar h ctxt t =
+ let
+ val congs = get_fundef_congs ctxt
+ val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) (* FIXME: Save in theory *)
+
+ fun matchcall (a $ b) = if a = Free fvar then SOME b else NONE
+ | matchcall _ = NONE
+
+ fun mk_tree' ctx t =
+ case matchcall t of
+ SOME arg => RCall (t, mk_tree' ctx arg)
+ | NONE =>
+ if not (exists_subterm (fn Free v => v = fvar | _ => false) t) then Leaf t
+ else
+ let val (r, dep, branches) = find_cong_rule ctx fvar h congs_deps t in
+ Cong (r, dep,
+ map (fn (ctx', fixes, assumes, st) =>
+ ((fixes, map (assume o cterm_of (ProofContext.theory_of ctx)) assumes),
+ mk_tree' ctx' st)) branches)
+ end
+ in
+ mk_tree' ctxt t
+ end
+
+
+fun inst_tree thy fvar f tr =
+ let
+ val cfvar = cterm_of thy fvar
+ val cf = cterm_of thy f
+
+ fun inst_term t =
+ subst_bound(f, abstract_over (fvar, t))
+
+ val inst_thm = forall_elim cf o forall_intr cfvar
+
+ fun inst_tree_aux (Leaf t) = Leaf t
+ | inst_tree_aux (Cong (crule, deps, branches)) =
+ Cong (inst_thm crule, deps, map inst_branch branches)
+ | inst_tree_aux (RCall (t, str)) =
+ RCall (inst_term t, inst_tree_aux str)
+ and inst_branch ((fxs, assms), str) =
+ ((fxs, map (assume o cterm_of thy o inst_term o prop_of) assms), inst_tree_aux str)
+ in
+ inst_tree_aux tr
+ end
+
+
+(* Poor man's contexts: Only fixes and assumes *)
+fun compose (fs1, as1) (fs2, as2) = (fs1 @ fs2, as1 @ as2)
+
+fun export_term (fixes, assumes) =
+ fold_rev (curry Logic.mk_implies o prop_of) assumes
+ #> fold_rev (Logic.all o Free) fixes
+
+fun export_thm thy (fixes, assumes) =
+ fold_rev (implies_intr o cprop_of) assumes
+ #> fold_rev (forall_intr o cterm_of thy o Free) fixes
+
+fun import_thm thy (fixes, athms) =
+ fold (forall_elim o cterm_of thy o Free) fixes
+ #> fold Thm.elim_implies athms
+
+
+(* folds in the order of the dependencies of a graph. *)
+fun fold_deps G f x =
+ let
+ fun fill_table i (T, x) =
+ case Inttab.lookup T i of
+ SOME _ => (T, x)
+ | NONE =>
+ let
+ val (T', x') = fold fill_table (IntGraph.imm_succs G i) (T, x)
+ val (v, x'') = f (the o Inttab.lookup T') i x'
+ in
+ (Inttab.update (i, v) T', x'')
+ end
+
+ val (T, x) = fold fill_table (IntGraph.keys G) (Inttab.empty, x)
+ in
+ (Inttab.fold (cons o snd) T [], x)
+ end
+
+fun traverse_tree rcOp tr =
+ let
+ fun traverse_help ctx (Leaf _) _ x = ([], x)
+ | traverse_help ctx (RCall (t, st)) u x =
+ rcOp ctx t u (traverse_help ctx st u x)
+ | traverse_help ctx (Cong (_, deps, branches)) u x =
+ let
+ fun sub_step lu i x =
+ let
+ val (ctx', subtree) = nth branches i
+ val used = fold_rev (append o lu) (IntGraph.imm_succs deps i) u
+ val (subs, x') = traverse_help (compose ctx ctx') subtree used x
+ val exported_subs = map (apfst (compose ctx')) subs (* FIXME: Right order of composition? *)
+ in
+ (exported_subs, x')
+ end
+ in
+ fold_deps deps sub_step x
+ |> apfst flat
+ end
+ in
+ snd o traverse_help ([], []) tr []
+ end
+
+fun rewrite_by_tree thy h ih x tr =
+ let
+ fun rewrite_help _ _ x (Leaf t) = (reflexive (cterm_of thy t), x)
+ | rewrite_help fix h_as x (RCall (_ $ arg, st)) =
+ let
+ val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *)
+
+ val iha = import_thm thy (fix, h_as) ha (* (a', h a') : G *)
+ |> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner))))
+ (* (a, h a) : G *)
+ val inst_ih = instantiate' [] [SOME (cterm_of thy arg)] ih
+ val eq = implies_elim (implies_elim inst_ih lRi) iha (* h a = f a *)
+
+ val h_a'_eq_h_a = combination (reflexive (cterm_of thy h)) inner
+ val h_a_eq_f_a = eq RS eq_reflection
+ val result = transitive h_a'_eq_h_a h_a_eq_f_a
+ in
+ (result, x')
+ end
+ | rewrite_help fix h_as x (Cong (crule, deps, branches)) =
+ let
+ fun sub_step lu i x =
+ let
+ val ((fixes, assumes), st) = nth branches i
+ val used = map lu (IntGraph.imm_succs deps i)
+ |> map (fn u_eq => (u_eq RS sym) RS eq_reflection)
+ |> filter_out Thm.is_reflexive
+
+ val assumes' = map (simplify (HOL_basic_ss addsimps used)) assumes
+
+ val (subeq, x') = rewrite_help (fix @ fixes) (h_as @ assumes') x st
+ val subeq_exp = export_thm thy (fixes, assumes) (subeq RS meta_eq_to_obj_eq)
+ in
+ (subeq_exp, x')
+ end
+
+ val (subthms, x') = fold_deps deps sub_step x
+ in
+ (fold_rev (curry op COMP) subthms crule, x')
+ end
+ in
+ rewrite_help [] [] x tr
+ end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/decompose.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,105 @@
+(* Title: HOL/Tools/Function/decompose.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Graph decomposition using "Shallow Dependency Pairs".
+*)
+
+signature DECOMPOSE =
+sig
+
+ val derive_chains : Proof.context -> tactic
+ -> (Termination.data -> int -> tactic)
+ -> Termination.data -> int -> tactic
+
+ val decompose_tac : Proof.context -> tactic
+ -> Termination.ttac
+
+end
+
+structure Decompose : DECOMPOSE =
+struct
+
+structure TermGraph = GraphFun(type key = term val ord = TermOrd.fast_term_ord);
+
+
+fun derive_chains ctxt chain_tac cont D = Termination.CALLS (fn (cs, i) =>
+ let
+ val thy = ProofContext.theory_of ctxt
+
+ fun prove_chain c1 c2 D =
+ if is_some (Termination.get_chain D c1 c2) then D else
+ let
+ val goal = HOLogic.mk_eq (HOLogic.mk_binop @{const_name "Relation.rel_comp"} (c1, c2),
+ Const (@{const_name Set.empty}, fastype_of c1))
+ |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
+
+ val chain = case FundefLib.try_proof (cterm_of thy goal) chain_tac of
+ FundefLib.Solved thm => SOME thm
+ | _ => NONE
+ in
+ Termination.note_chain c1 c2 chain D
+ end
+ in
+ cont (fold_product prove_chain cs cs D) i
+ end)
+
+
+fun mk_dgraph D cs =
+ TermGraph.empty
+ |> fold (fn c => TermGraph.new_node (c,())) cs
+ |> fold_product (fn c1 => fn c2 =>
+ if is_none (Termination.get_chain D c1 c2 |> the_default NONE)
+ then TermGraph.add_edge (c1, c2) else I)
+ cs cs
+
+
+fun ucomp_empty_tac T =
+ REPEAT_ALL_NEW (rtac @{thm union_comp_emptyR}
+ ORELSE' rtac @{thm union_comp_emptyL}
+ ORELSE' SUBGOAL (fn (_ $ (_ $ (_ $ c1 $ c2) $ _), i) => rtac (T c1 c2) i))
+
+fun regroup_calls_tac cs = Termination.CALLS (fn (cs', i) =>
+ let
+ val is = map (fn c => find_index (curry op aconv c) cs') cs
+ in
+ CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv is))) i
+ end)
+
+
+fun solve_trivial_tac D = Termination.CALLS
+(fn ([c], i) =>
+ (case Termination.get_chain D c c of
+ SOME (SOME thm) => rtac @{thm wf_no_loop} i
+ THEN rtac thm i
+ | _ => no_tac)
+ | _ => no_tac)
+
+fun decompose_tac' ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
+ let
+ val G = mk_dgraph D cs
+ val sccs = TermGraph.strong_conn G
+
+ fun split [SCC] i = (solve_trivial_tac D i ORELSE cont D i)
+ | split (SCC::rest) i =
+ regroup_calls_tac SCC i
+ THEN rtac @{thm wf_union_compatible} i
+ THEN rtac @{thm less_by_empty} (i + 2)
+ THEN ucomp_empty_tac (the o the oo Termination.get_chain D) (i + 2)
+ THEN split rest (i + 1)
+ THEN (solve_trivial_tac D i ORELSE cont D i)
+ in
+ if length sccs > 1 then split sccs i
+ else solve_trivial_tac D i ORELSE err_cont D i
+ end)
+
+fun decompose_tac ctxt chain_tac cont err_cont =
+ derive_chains ctxt chain_tac
+ (decompose_tac' ctxt cont err_cont)
+
+fun auto_decompose_tac ctxt =
+ Termination.TERMINATION ctxt
+ (decompose_tac ctxt (auto_tac (local_clasimpset_of ctxt))
+ (K (K all_tac)) (K (K no_tac)))
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/descent.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,44 @@
+(* Title: HOL/Tools/Function/descent.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Descent proofs for termination
+*)
+
+
+signature DESCENT =
+sig
+
+ val derive_diag : Proof.context -> tactic -> (Termination.data -> int -> tactic)
+ -> Termination.data -> int -> tactic
+
+ val derive_all : Proof.context -> tactic -> (Termination.data -> int -> tactic)
+ -> Termination.data -> int -> tactic
+
+end
+
+
+structure Descent : DESCENT =
+struct
+
+fun gen_descent diag ctxt tac cont D = Termination.CALLS (fn (cs, i) =>
+ let
+ val thy = ProofContext.theory_of ctxt
+ val measures_of = Termination.get_measures D
+
+ fun derive c D =
+ let
+ val (_, p, _, q, _, _) = Termination.dest_call D c
+ in
+ if diag andalso p = q
+ then fold (fn m => Termination.derive_descent thy tac c m m) (measures_of p) D
+ else fold_product (Termination.derive_descent thy tac c)
+ (measures_of p) (measures_of q) D
+ end
+ in
+ cont (FundefCommon.PROFILE "deriving descents" (fold derive cs) D) i
+ end)
+
+val derive_diag = gen_descent true
+val derive_all = gen_descent false
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fundef.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,226 @@
+(* Title: HOL/Tools/Function/fundef.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Isar commands.
+*)
+
+signature FUNDEF =
+sig
+ val add_fundef : (binding * typ option * mixfix) list
+ -> (Attrib.binding * term) list
+ -> FundefCommon.fundef_config
+ -> local_theory
+ -> Proof.state
+ val add_fundef_cmd : (binding * string option * mixfix) list
+ -> (Attrib.binding * string) list
+ -> FundefCommon.fundef_config
+ -> local_theory
+ -> Proof.state
+
+ val termination_proof : term option -> local_theory -> Proof.state
+ val termination_proof_cmd : string option -> local_theory -> Proof.state
+ val termination : term option -> local_theory -> Proof.state
+ val termination_cmd : string option -> local_theory -> Proof.state
+
+ val setup : theory -> theory
+ val get_congs : Proof.context -> thm list
+end
+
+
+structure Fundef : FUNDEF =
+struct
+
+open FundefLib
+open FundefCommon
+
+val simp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Code.add_default_eqn_attribute,
+ Nitpick_Const_Simp_Thms.add,
+ Quickcheck_RecFun_Simp_Thms.add]
+
+val psimp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Nitpick_Const_Psimp_Thms.add]
+
+fun note_theorem ((name, atts), ths) =
+ LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
+
+fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
+
+fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
+ let
+ val spec = post simps
+ |> map (apfst (apsnd (fn ats => moreatts @ ats)))
+ |> map (apfst (apfst extra_qualify))
+
+ val (saved_spec_simps, lthy) =
+ fold_map (LocalTheory.note Thm.generatedK) spec lthy
+
+ val saved_simps = flat (map snd saved_spec_simps)
+ val simps_by_f = sort saved_simps
+
+ fun add_for_f fname simps =
+ note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
+ in
+ (saved_simps,
+ fold2 add_for_f fnames simps_by_f lthy)
+ end
+
+fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
+ let
+ val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
+ val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
+ val fixes = map (apfst (apfst Binding.name_of)) fixes0;
+ val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
+ val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
+
+ val defname = mk_defname fixes
+
+ val ((goalstate, cont), lthy) =
+ FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
+
+ fun afterqed [[proof]] lthy =
+ let
+ val FundefResult {fs, R, psimps, trsimps, simple_pinducts, termination,
+ domintros, cases, ...} =
+ cont (Thm.close_derivation proof)
+
+ val fnames = map (fst o fst) fixes
+ val qualify = Long_Name.qualify defname
+ val addsmps = add_simps fnames post sort_cont
+
+ val (((psimps', pinducts'), (_, [termination'])), lthy) =
+ lthy
+ |> addsmps (Binding.qualify false "partial") "psimps"
+ psimp_attribs psimps
+ ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
+ ||>> note_theorem ((qualify "pinduct",
+ [Attrib.internal (K (RuleCases.case_names cnames)),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
+ ||>> note_theorem ((qualify "termination", []), [termination])
+ ||> (snd o note_theorem ((qualify "cases",
+ [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
+ ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
+
+ val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
+ pinducts=snd pinducts', termination=termination',
+ fs=fs, R=R, defname=defname }
+ val _ =
+ if not is_external then ()
+ else Specification.print_consts lthy (K false) (map fst fixes)
+ in
+ lthy
+ |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
+ end
+ in
+ lthy
+ |> is_external ? LocalTheory.set_group (serial_string ())
+ |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+ |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
+ end
+
+val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
+val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
+
+fun gen_termination_proof prep_term raw_term_opt lthy =
+ let
+ val term_opt = Option.map (prep_term lthy) raw_term_opt
+ val data = the (case term_opt of
+ SOME t => (import_fundef_data t lthy
+ handle Option.Option =>
+ error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
+ | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
+
+ val FundefCtxData { termination, R, add_simps, case_names, psimps,
+ pinducts, defname, ...} = data
+ val domT = domain_type (fastype_of R)
+ val goal = HOLogic.mk_Trueprop
+ (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
+ fun afterqed [[totality]] lthy =
+ let
+ val totality = Thm.close_derivation totality
+ val remove_domain_condition =
+ full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
+ val tsimps = map remove_domain_condition psimps
+ val tinduct = map remove_domain_condition pinducts
+ val qualify = Long_Name.qualify defname;
+ in
+ lthy
+ |> add_simps I "simps" simp_attribs tsimps |> snd
+ |> note_theorem
+ ((qualify "induct",
+ [Attrib.internal (K (RuleCases.case_names case_names))]),
+ tinduct) |> snd
+ end
+ in
+ lthy
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
+ [([Goal.norm_result termination], [])])] |> snd
+ |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+ end
+
+val termination_proof = gen_termination_proof Syntax.check_term;
+val termination_proof_cmd = gen_termination_proof Syntax.read_term;
+
+fun termination term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof term_opt;
+
+fun termination_cmd term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof_cmd term_opt;
+
+
+(* Datatype hook to declare datatype congs as "fundef_congs" *)
+
+
+fun add_case_cong n thy =
+ Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
+ (Datatype.get_datatype thy n |> the
+ |> #case_cong
+ |> safe_mk_meta_eq)))
+ thy
+
+val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
+
+
+(* setup *)
+
+val setup =
+ Attrib.setup @{binding fundef_cong}
+ (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
+ "declaration of congruence rule for function definitions"
+ #> setup_case_cong
+ #> FundefRelation.setup
+ #> FundefCommon.TerminationSimps.setup
+
+val get_congs = FundefCtxTree.get_fundef_congs
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+ OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
+ (fundef_parser default_config
+ >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
+
+val _ =
+ OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
+ (Scan.option P.term >> termination_cmd);
+
+end;
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fundef_common.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,343 @@
+(* Title: HOL/Tools/Function/fundef_common.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Common definitions and other infrastructure.
+*)
+
+structure FundefCommon =
+struct
+
+local open FundefLib in
+
+(* Profiling *)
+val profile = ref false;
+
+fun PROFILE msg = if !profile then timeap_msg msg else I
+
+
+val acc_const_name = @{const_name "accp"}
+fun mk_acc domT R =
+ Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R
+
+val function_name = suffix "C"
+val graph_name = suffix "_graph"
+val rel_name = suffix "_rel"
+val dom_name = suffix "_dom"
+
+(* Termination rules *)
+
+structure TerminationRule = GenericDataFun
+(
+ type T = thm list
+ val empty = []
+ val extend = I
+ fun merge _ = Thm.merge_thms
+);
+
+val get_termination_rules = TerminationRule.get
+val store_termination_rule = TerminationRule.map o cons
+val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
+
+
+(* Function definition result data *)
+
+datatype fundef_result =
+ FundefResult of
+ {
+ fs: term list,
+ G: term,
+ R: term,
+
+ psimps : thm list,
+ trsimps : thm list option,
+
+ simple_pinducts : thm list,
+ cases : thm,
+ termination : thm,
+ domintros : thm list option
+ }
+
+
+datatype fundef_context_data =
+ FundefCtxData of
+ {
+ defname : string,
+
+ (* contains no logical entities: invariant under morphisms *)
+ add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list
+ -> local_theory -> thm list * local_theory,
+ case_names : string list,
+
+ fs : term list,
+ R : term,
+
+ psimps: thm list,
+ pinducts: thm list,
+ termination: thm
+ }
+
+fun morph_fundef_data (FundefCtxData {add_simps, case_names, fs, R,
+ psimps, pinducts, termination, defname}) phi =
+ let
+ val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
+ val name = Binding.name_of o Morphism.binding phi o Binding.name
+ in
+ FundefCtxData { add_simps = add_simps, case_names = case_names,
+ fs = map term fs, R = term R, psimps = fact psimps,
+ pinducts = fact pinducts, termination = thm termination,
+ defname = name defname }
+ end
+
+structure FundefData = GenericDataFun
+(
+ type T = (term * fundef_context_data) Item_Net.T;
+ val empty = Item_Net.init
+ (op aconv o pairself fst : (term * fundef_context_data) * (term * fundef_context_data) -> bool)
+ fst;
+ val copy = I;
+ val extend = I;
+ fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
+);
+
+val get_fundef = FundefData.get o Context.Proof;
+
+
+(* Generally useful?? *)
+fun lift_morphism thy f =
+ let
+ val term = Drule.term_rule thy f
+ in
+ Morphism.thm_morphism f $> Morphism.term_morphism term
+ $> Morphism.typ_morphism (Logic.type_map term)
+ end
+
+fun import_fundef_data t ctxt =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val ct = cterm_of thy t
+ val inst_morph = lift_morphism thy o Thm.instantiate
+
+ fun match (trm, data) =
+ SOME (morph_fundef_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
+ handle Pattern.MATCH => NONE
+ in
+ get_first match (Item_Net.retrieve (get_fundef ctxt) t)
+ end
+
+fun import_last_fundef ctxt =
+ case Item_Net.content (get_fundef ctxt) of
+ [] => NONE
+ | (t, data) :: _ =>
+ let
+ val ([t'], ctxt') = Variable.import_terms true [t] ctxt
+ in
+ import_fundef_data t' ctxt'
+ end
+
+val all_fundef_data = Item_Net.content o get_fundef
+
+fun add_fundef_data (data as FundefCtxData {fs, termination, ...}) =
+ FundefData.map (fold (fn f => Item_Net.insert (f, data)) fs)
+ #> store_termination_rule termination
+
+
+(* Simp rules for termination proofs *)
+
+structure TerminationSimps = NamedThmsFun
+(
+ val name = "termination_simp"
+ val description = "Simplification rule for termination proofs"
+);
+
+
+(* Default Termination Prover *)
+
+structure TerminationProver = GenericDataFun
+(
+ type T = Proof.context -> Proof.method
+ val empty = (fn _ => error "Termination prover not configured")
+ val extend = I
+ fun merge _ (a,b) = b (* FIXME *)
+);
+
+val set_termination_prover = TerminationProver.put
+val get_termination_prover = TerminationProver.get o Context.Proof
+
+
+(* Configuration management *)
+datatype fundef_opt
+ = Sequential
+ | Default of string
+ | DomIntros
+ | Tailrec
+
+datatype fundef_config
+ = FundefConfig of
+ {
+ sequential: bool,
+ default: string,
+ domintros: bool,
+ tailrec: bool
+ }
+
+fun apply_opt Sequential (FundefConfig {sequential, default, domintros,tailrec}) =
+ FundefConfig {sequential=true, default=default, domintros=domintros, tailrec=tailrec}
+ | apply_opt (Default d) (FundefConfig {sequential, default, domintros,tailrec}) =
+ FundefConfig {sequential=sequential, default=d, domintros=domintros, tailrec=tailrec}
+ | apply_opt DomIntros (FundefConfig {sequential, default, domintros,tailrec}) =
+ FundefConfig {sequential=sequential, default=default, domintros=true,tailrec=tailrec}
+ | apply_opt Tailrec (FundefConfig {sequential, default, domintros,tailrec}) =
+ FundefConfig {sequential=sequential, default=default, domintros=domintros,tailrec=true}
+
+val default_config =
+ FundefConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*),
+ domintros=false, tailrec=false }
+
+
+(* Analyzing function equations *)
+
+fun split_def ctxt geq =
+ let
+ fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
+ val qs = Term.strip_qnt_vars "all" geq
+ val imp = Term.strip_qnt_body "all" geq
+ val (gs, eq) = Logic.strip_horn imp
+
+ val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
+ handle TERM _ => error (input_error "Not an equation")
+
+ val (head, args) = strip_comb f_args
+
+ val fname = fst (dest_Free head)
+ handle TERM _ => error (input_error "Head symbol must not be a bound variable")
+ in
+ (fname, qs, gs, args, rhs)
+ end
+
+(* Check for all sorts of errors in the input *)
+fun check_defs ctxt fixes eqs =
+ let
+ val fnames = map (fst o fst) fixes
+
+ fun check geq =
+ let
+ fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
+
+ val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
+
+ val _ = fname mem fnames
+ orelse input_error
+ ("Head symbol of left hand side must be "
+ ^ plural "" "one out of " fnames ^ commas_quote fnames)
+
+ val _ = length args > 0 orelse input_error "Function has no arguments:"
+
+ fun add_bvs t is = add_loose_bnos (t, 0, is)
+ val rvs = (add_bvs rhs [] \\ fold add_bvs args [])
+ |> map (fst o nth (rev qs))
+
+ val _ = null rvs orelse input_error
+ ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
+ ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
+
+ val _ = forall (not o Term.exists_subterm
+ (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
+ orelse input_error "Defined function may not occur in premises or arguments"
+
+ val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
+ val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
+ val _ = null funvars
+ orelse (warning (cat_lines
+ ["Bound variable" ^ plural " " "s " funvars
+ ^ commas_quote (map fst funvars) ^
+ " occur" ^ plural "s" "" funvars ^ " in function position.",
+ "Misspelled constructor???"]); true)
+ in
+ (fname, length args)
+ end
+
+ val _ = AList.group (op =) (map check eqs)
+ |> map (fn (fname, ars) =>
+ length (distinct (op =) ars) = 1
+ orelse error ("Function " ^ quote fname ^
+ " has different numbers of arguments in different equations"))
+
+ fun check_sorts ((fname, fT), _) =
+ Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
+ orelse error (cat_lines
+ ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
+ setmp show_sorts true (Syntax.string_of_typ ctxt) fT])
+
+ val _ = map check_sorts fixes
+ in
+ ()
+ end
+
+(* Preprocessors *)
+
+type fixes = ((string * typ) * mixfix) list
+type 'a spec = (Attrib.binding * 'a list) list
+type preproc = fundef_config -> Proof.context -> fixes -> term spec
+ -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
+
+val fname_of = fst o dest_Free o fst o strip_comb o fst
+ o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
+
+fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
+ | mk_case_names _ n 0 = []
+ | mk_case_names _ n 1 = [n]
+ | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
+
+fun empty_preproc check _ ctxt fixes spec =
+ let
+ val (bnds, tss) = split_list spec
+ val ts = flat tss
+ val _ = check ctxt fixes ts
+ val fnames = map (fst o fst) fixes
+ val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
+
+ fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1)
+ (indices ~~ xs)
+ |> map (map snd)
+
+ (* using theorem names for case name currently disabled *)
+ val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
+ in
+ (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
+ end
+
+structure Preprocessor = GenericDataFun
+(
+ type T = preproc
+ val empty : T = empty_preproc check_defs
+ val extend = I
+ fun merge _ (a, _) = a
+);
+
+val get_preproc = Preprocessor.get o Context.Proof
+val set_preproc = Preprocessor.map o K
+
+
+
+local
+ structure P = OuterParse and K = OuterKeyword
+
+ val option_parser =
+ P.group "option" ((P.reserved "sequential" >> K Sequential)
+ || ((P.reserved "default" |-- P.term) >> Default)
+ || (P.reserved "domintros" >> K DomIntros)
+ || (P.reserved "tailrec" >> K Tailrec))
+
+ fun config_parser default =
+ (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
+ >> (fn opts => fold apply_opt opts default)
+in
+ fun fundef_parser default_cfg =
+ config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
+end
+
+
+end
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fundef_core.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,954 @@
+(* Title: HOL/Tools/Function/fundef_core.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions:
+Main functionality.
+*)
+
+signature FUNDEF_CORE =
+sig
+ val prepare_fundef : FundefCommon.fundef_config
+ -> string (* defname *)
+ -> ((bstring * typ) * mixfix) list (* defined symbol *)
+ -> ((bstring * typ) list * term list * term * term) list (* specification *)
+ -> local_theory
+
+ -> (term (* f *)
+ * thm (* goalstate *)
+ * (thm -> FundefCommon.fundef_result) (* continuation *)
+ ) * local_theory
+
+end
+
+structure FundefCore : FUNDEF_CORE =
+struct
+
+val boolT = HOLogic.boolT
+val mk_eq = HOLogic.mk_eq
+
+open FundefLib
+open FundefCommon
+
+datatype globals =
+ Globals of {
+ fvar: term,
+ domT: typ,
+ ranT: typ,
+ h: term,
+ y: term,
+ x: term,
+ z: term,
+ a: term,
+ P: term,
+ D: term,
+ Pbool:term
+}
+
+
+datatype rec_call_info =
+ RCInfo of
+ {
+ RIvs: (string * typ) list, (* Call context: fixes and assumes *)
+ CCas: thm list,
+ rcarg: term, (* The recursive argument *)
+
+ llRI: thm,
+ h_assum: term
+ }
+
+
+datatype clause_context =
+ ClauseContext of
+ {
+ ctxt : Proof.context,
+
+ qs : term list,
+ gs : term list,
+ lhs: term,
+ rhs: term,
+
+ cqs: cterm list,
+ ags: thm list,
+ case_hyp : thm
+ }
+
+
+fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
+ ClauseContext { ctxt = ProofContext.transfer thy ctxt,
+ qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
+
+
+datatype clause_info =
+ ClauseInfo of
+ {
+ no: int,
+ qglr : ((string * typ) list * term list * term * term),
+ cdata : clause_context,
+
+ tree: FundefCtxTree.ctx_tree,
+ lGI: thm,
+ RCs: rec_call_info list
+ }
+
+
+(* Theory dependencies. *)
+val Pair_inject = @{thm Product_Type.Pair_inject};
+
+val acc_induct_rule = @{thm accp_induct_rule};
+
+val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
+val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
+val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
+
+val acc_downward = @{thm accp_downward};
+val accI = @{thm accp.accI};
+val case_split = @{thm HOL.case_split};
+val fundef_default_value = @{thm FunDef.fundef_default_value};
+val not_acc_down = @{thm not_accp_down};
+
+
+
+fun find_calls tree =
+ let
+ fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
+ | add_Ri _ _ _ _ = raise Match
+ in
+ rev (FundefCtxTree.traverse_tree add_Ri tree [])
+ end
+
+
+(** building proof obligations *)
+
+fun mk_compat_proof_obligations domT ranT fvar f glrs =
+ let
+ fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
+ let
+ val shift = incr_boundvars (length qs')
+ in
+ Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
+ HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
+ |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
+ |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
+ |> curry abstract_over fvar
+ |> curry subst_bound f
+ end
+ in
+ map mk_impl (unordered_pairs glrs)
+ end
+
+
+fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
+ let
+ fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
+ HOLogic.mk_Trueprop Pbool
+ |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ in
+ HOLogic.mk_Trueprop Pbool
+ |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
+ |> mk_forall_rename ("x", x)
+ |> mk_forall_rename ("P", Pbool)
+ end
+
+(** making a context with it's own local bindings **)
+
+fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
+ let
+ val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
+ |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
+
+ val thy = ProofContext.theory_of ctxt'
+
+ fun inst t = subst_bounds (rev qs, t)
+ val gs = map inst pre_gs
+ val lhs = inst pre_lhs
+ val rhs = inst pre_rhs
+
+ val cqs = map (cterm_of thy) qs
+ val ags = map (assume o cterm_of thy) gs
+
+ val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+ in
+ ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
+ cqs = cqs, ags = ags, case_hyp = case_hyp }
+ end
+
+
+(* lowlevel term function *)
+fun abstract_over_list vs body =
+ let
+ exception SAME;
+ fun abs lev v tm =
+ if v aconv tm then Bound lev
+ else
+ (case tm of
+ Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
+ | t $ u => (abs lev v t $ (abs lev v u handle SAME => u) handle SAME => t $ abs lev v u)
+ | _ => raise SAME);
+ in
+ fold_index (fn (i,v) => fn t => abs i v t handle SAME => t) vs body
+ end
+
+
+
+fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
+ let
+ val Globals {h, fvar, x, ...} = globals
+
+ val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
+ val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+
+ (* Instantiate the GIntro thm with "f" and import into the clause context. *)
+ val lGI = GIntro_thm
+ |> forall_elim (cert f)
+ |> fold forall_elim cqs
+ |> fold Thm.elim_implies ags
+
+ fun mk_call_info (rcfix, rcassm, rcarg) RI =
+ let
+ val llRI = RI
+ |> fold forall_elim cqs
+ |> fold (forall_elim o cert o Free) rcfix
+ |> fold Thm.elim_implies ags
+ |> fold Thm.elim_implies rcassm
+
+ val h_assum =
+ HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (Logic.all o Free) rcfix
+ |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
+ |> abstract_over_list (rev qs)
+ in
+ RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
+ end
+
+ val RC_infos = map2 mk_call_info RCs RIntro_thms
+ in
+ ClauseInfo
+ {
+ no=no,
+ cdata=cdata,
+ qglr=qglr,
+
+ lGI=lGI,
+ RCs=RC_infos,
+ tree=tree
+ }
+ end
+
+
+
+
+
+
+
+(* replace this by a table later*)
+fun store_compat_thms 0 thms = []
+ | store_compat_thms n thms =
+ let
+ val (thms1, thms2) = chop n thms
+ in
+ (thms1 :: store_compat_thms (n - 1) thms2)
+ end
+
+(* expects i <= j *)
+fun lookup_compat_thm i j cts =
+ nth (nth cts (i - 1)) (j - i)
+
+(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
+(* if j < i, then turn around *)
+fun get_compat_thm thy cts i j ctxi ctxj =
+ let
+ val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
+ val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
+
+ val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
+ in if j < i then
+ let
+ val compat = lookup_compat_thm j i cts
+ in
+ compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ |> fold Thm.elim_implies agsj
+ |> fold Thm.elim_implies agsi
+ |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
+ end
+ else
+ let
+ val compat = lookup_compat_thm i j cts
+ in
+ compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ |> fold Thm.elim_implies agsi
+ |> fold Thm.elim_implies agsj
+ |> Thm.elim_implies (assume lhsi_eq_lhsj)
+ |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
+ end
+ end
+
+
+
+
+(* Generates the replacement lemma in fully quantified form. *)
+fun mk_replacement_lemma thy h ih_elim clause =
+ let
+ val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
+ local open Conv in
+ val ih_conv = arg1_conv o arg_conv o arg_conv
+ end
+
+ val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
+
+ val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
+ val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+
+ val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
+
+ val replace_lemma = (eql RS meta_eq_to_obj_eq)
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) h_assums
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+ |> Thm.close_derivation
+ in
+ replace_lemma
+ end
+
+
+fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
+ let
+ val Globals {h, y, x, fvar, ...} = globals
+ val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
+ val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
+
+ val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
+ = mk_clause_context x ctxti cdescj
+
+ val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
+ val compat = get_compat_thm thy compat_store i j cctxi cctxj
+ val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+
+ val RLj_import =
+ RLj |> fold forall_elim cqsj'
+ |> fold Thm.elim_implies agsj'
+ |> fold Thm.elim_implies Ghsj'
+
+ val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+ val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
+ in
+ (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
+ |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
+ |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
+ |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
+ |> fold_rev (implies_intr o cprop_of) Ghsj'
+ |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
+ |> implies_intr (cprop_of y_eq_rhsj'h)
+ |> implies_intr (cprop_of lhsi_eq_lhsj')
+ |> fold_rev forall_intr (cterm_of thy h :: cqsj')
+ end
+
+
+
+fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
+ let
+ val Globals {x, y, ranT, fvar, ...} = globals
+ val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
+ val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+ val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
+
+ fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
+ |> fold_rev (implies_intr o cprop_of) CCas
+ |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+ val existence = fold (curry op COMP o prep_RC) RCs lGI
+
+ val P = cterm_of thy (mk_eq (y, rhsC))
+ val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+
+ val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
+
+ val uniqueness = G_cases
+ |> forall_elim (cterm_of thy lhs)
+ |> forall_elim (cterm_of thy y)
+ |> forall_elim P
+ |> Thm.elim_implies G_lhs_y
+ |> fold Thm.elim_implies unique_clauses
+ |> implies_intr (cprop_of G_lhs_y)
+ |> forall_intr (cterm_of thy y)
+
+ val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
+
+ val exactly_one =
+ ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
+ |> curry (op COMP) existence
+ |> curry (op COMP) uniqueness
+ |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+
+ val function_value =
+ existence
+ |> implies_intr ihyp
+ |> implies_intr (cprop_of case_hyp)
+ |> forall_intr (cterm_of thy x)
+ |> forall_elim (cterm_of thy lhs)
+ |> curry (op RS) refl
+ in
+ (exactly_one, function_value)
+ end
+
+
+
+
+fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
+ let
+ val Globals {h, domT, ranT, x, ...} = globals
+ val thy = ProofContext.theory_of ctxt
+
+ (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
+ Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
+ |> cterm_of thy
+
+ val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
+ val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
+ val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
+ |> instantiate' [] [NONE, SOME (cterm_of thy h)]
+
+ val _ = Output.debug (K "Proving Replacement lemmas...")
+ val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
+
+ val _ = Output.debug (K "Proving cases for unique existence...")
+ val (ex1s, values) =
+ split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
+
+ val _ = Output.debug (K "Proving: Graph is a function")
+ val graph_is_function = complete
+ |> Thm.forall_elim_vars 0
+ |> fold (curry op COMP) ex1s
+ |> implies_intr (ihyp)
+ |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+ |> forall_intr (cterm_of thy x)
+ |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
+ |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+
+ val goalstate = Conjunction.intr graph_is_function complete
+ |> Thm.close_derivation
+ |> Goal.protect
+ |> fold_rev (implies_intr o cprop_of) compat
+ |> implies_intr (cprop_of complete)
+ in
+ (goalstate, values)
+ end
+
+
+fun define_graph Gname fvar domT ranT clauses RCss lthy =
+ let
+ val GT = domT --> ranT --> boolT
+ val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
+
+ fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
+ let
+ fun mk_h_assm (rcfix, rcassm, rcarg) =
+ HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (Logic.all o Free) rcfix
+ in
+ HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
+ |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev Logic.all (fvar :: qs)
+ end
+
+ val G_intros = map2 mk_GIntro clauses RCss
+
+ val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
+ FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
+ in
+ ((G, GIntro_thms, G_elim, G_induct), lthy)
+ end
+
+
+
+fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
+ let
+ val f_def =
+ Abs ("x", domT, Const ("FunDef.THE_default", ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
+ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
+ |> Syntax.check_term lthy
+
+ val ((f, (_, f_defthm)), lthy) =
+ LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
+ in
+ ((f, f_defthm), lthy)
+ end
+
+
+fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
+ let
+
+ val RT = domT --> domT --> boolT
+ val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
+
+ fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
+ HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev (Logic.all o Free) rcfix
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
+
+ val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
+
+ val (RIntro_thmss, (R, R_elim, _, lthy)) =
+ fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
+ in
+ ((R, RIntro_thmss, R_elim), lthy)
+ end
+
+
+fun fix_globals domT ranT fvar ctxt =
+ let
+ val ([h, y, x, z, a, D, P, Pbool],ctxt') =
+ Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
+ in
+ (Globals {h = Free (h, domT --> ranT),
+ y = Free (y, ranT),
+ x = Free (x, domT),
+ z = Free (z, domT),
+ a = Free (a, domT),
+ D = Free (D, domT --> boolT),
+ P = Free (P, domT --> boolT),
+ Pbool = Free (Pbool, boolT),
+ fvar = fvar,
+ domT = domT,
+ ranT = ranT
+ },
+ ctxt')
+ end
+
+
+
+fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
+ let
+ fun inst_term t = subst_bound(f, abstract_over (fvar, t))
+ in
+ (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+ end
+
+
+
+(**********************************************************
+ * PROVING THE RULES
+ **********************************************************)
+
+fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
+ let
+ val Globals {domT, z, ...} = globals
+
+ fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
+ let
+ val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
+ val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
+ in
+ ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
+ |> (fn it => it COMP graph_is_function)
+ |> implies_intr z_smaller
+ |> forall_intr (cterm_of thy z)
+ |> (fn it => it COMP valthm)
+ |> implies_intr lhs_acc
+ |> asm_simplify (HOL_basic_ss addsimps [f_iff])
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+ in
+ map2 mk_psimp clauses valthms
+ end
+
+
+(** Induction rule **)
+
+
+val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
+
+
+fun binder_conv cv ctxt = Conv.arg_conv (Conv.abs_conv (K cv) ctxt);
+
+fun mk_partial_induct_rule thy globals R complete_thm clauses =
+ let
+ val Globals {domT, x, z, a, P, D, ...} = globals
+ val acc_R = mk_acc domT R
+
+ val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+ val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
+
+ val D_subset = cterm_of thy (Logic.all x
+ (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
+
+ val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
+ Logic.all x
+ (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
+ HOLogic.mk_Trueprop (D $ z)))))
+ |> cterm_of thy
+
+
+ (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (P $ Bound 0)))
+ |> cterm_of thy
+
+ val aihyp = assume ihyp
+
+ fun prove_case clause =
+ let
+ val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
+ qglr = (oqs, _, _, _), ...} = clause
+
+ val case_hyp_conv = K (case_hyp RS eq_reflection)
+ local open Conv in
+ val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
+ val sih = fconv_rule (binder_conv (arg1_conv (arg_conv (arg_conv case_hyp_conv))) ctxt) aihyp
+ end
+
+ fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
+ sih |> forall_elim (cterm_of thy rcarg)
+ |> Thm.elim_implies llRI
+ |> fold_rev (implies_intr o cprop_of) CCas
+ |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+ val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
+
+ val step = HOLogic.mk_Trueprop (P $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ |> cterm_of thy
+
+ val P_lhs = assume step
+ |> fold forall_elim cqs
+ |> Thm.elim_implies lhs_D
+ |> fold Thm.elim_implies ags
+ |> fold Thm.elim_implies P_recs
+
+ val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
+ |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
+ |> symmetric (* P lhs == P x *)
+ |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+ in
+ (res, step)
+ end
+
+ val (cases, steps) = split_list (map prove_case clauses)
+
+ val istep = complete_thm
+ |> Thm.forall_elim_vars 0
+ |> fold (curry op COMP) cases (* P x *)
+ |> implies_intr ihyp
+ |> implies_intr (cprop_of x_D)
+ |> forall_intr (cterm_of thy x)
+
+ val subset_induct_rule =
+ acc_subset_induct
+ |> (curry op COMP) (assume D_subset)
+ |> (curry op COMP) (assume D_dcl)
+ |> (curry op COMP) (assume a_D)
+ |> (curry op COMP) istep
+ |> fold_rev implies_intr steps
+ |> implies_intr a_D
+ |> implies_intr D_dcl
+ |> implies_intr D_subset
+
+ val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
+
+ val simple_induct_rule =
+ subset_induct_rule
+ |> forall_intr (cterm_of thy D)
+ |> forall_elim (cterm_of thy acc_R)
+ |> assume_tac 1 |> Seq.hd
+ |> (curry op COMP) (acc_downward
+ |> (instantiate' [SOME (ctyp_of thy domT)]
+ (map (SOME o cterm_of thy) [R, x, z]))
+ |> forall_intr (cterm_of thy z)
+ |> forall_intr (cterm_of thy x))
+ |> forall_intr (cterm_of thy a)
+ |> forall_intr (cterm_of thy P)
+ in
+ simple_induct_rule
+ end
+
+
+
+(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
+fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
+ qglr = (oqs, _, _, _), ...} = clause
+ val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> cterm_of thy
+ in
+ Goal.init goal
+ |> (SINGLE (resolve_tac [accI] 1)) |> the
+ |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
+ |> (SINGLE (auto_tac (local_clasimpset_of ctxt))) |> the
+ |> Goal.conclude
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+
+
+
+(** Termination rule **)
+
+val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
+val wf_in_rel = @{thm FunDef.wf_in_rel};
+val in_rel_def = @{thm FunDef.in_rel_def};
+
+fun mk_nest_term_case thy globals R' ihyp clause =
+ let
+ val Globals {x, z, ...} = globals
+ val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
+ qglr=(oqs, _, _, _), ...} = clause
+
+ val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
+
+ fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
+ let
+ val used = map (fn (ctx,thm) => FundefCtxTree.export_thm thy ctx thm) (u @ sub)
+
+ val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
+ |> FundefCtxTree.export_term (fixes, assumes)
+ |> fold_rev (curry Logic.mk_implies o prop_of) ags
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ |> cterm_of thy
+
+ val thm = assume hyp
+ |> fold forall_elim cqs
+ |> fold Thm.elim_implies ags
+ |> FundefCtxTree.import_thm thy (fixes, assumes)
+ |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
+
+ val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
+
+ val acc = thm COMP ih_case
+ val z_acc_local = acc
+ |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
+
+ val ethm = z_acc_local
+ |> FundefCtxTree.export_thm thy (fixes,
+ z_eq_arg :: case_hyp :: ags @ assumes)
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+
+ val sub' = sub @ [(([],[]), acc)]
+ in
+ (sub', (hyp :: hyps, ethm :: thms))
+ end
+ | step _ _ _ _ = raise Match
+ in
+ FundefCtxTree.traverse_tree step tree
+ end
+
+
+fun mk_nest_term_rule thy globals R R_cases clauses =
+ let
+ val Globals { domT, x, z, ... } = globals
+ val acc_R = mk_acc domT R
+
+ val R' = Free ("R", fastype_of R)
+
+ val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
+ val inrel_R = Const ("FunDef.in_rel", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
+
+ val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name "Wellfounded.wfP"}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
+
+ (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (acc_R $ Bound 0)))
+ |> cterm_of thy
+
+ val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
+
+ val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
+
+ val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
+ in
+ R_cases
+ |> forall_elim (cterm_of thy z)
+ |> forall_elim (cterm_of thy x)
+ |> forall_elim (cterm_of thy (acc_R $ z))
+ |> curry op COMP (assume R_z_x)
+ |> fold_rev (curry op COMP) cases
+ |> implies_intr R_z_x
+ |> forall_intr (cterm_of thy z)
+ |> (fn it => it COMP accI)
+ |> implies_intr ihyp
+ |> forall_intr (cterm_of thy x)
+ |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
+ |> curry op RS (assume wfR')
+ |> forall_intr_vars
+ |> (fn it => it COMP allI)
+ |> fold implies_intr hyps
+ |> implies_intr wfR'
+ |> forall_intr (cterm_of thy R')
+ |> forall_elim (cterm_of thy (inrel_R))
+ |> curry op RS wf_in_rel
+ |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
+ |> forall_intr (cterm_of thy Rrel)
+ end
+
+
+
+(* Tail recursion (probably very fragile)
+ *
+ * FIXME:
+ * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
+ * - Must we really replace the fvar by f here?
+ * - Splitting is not configured automatically: Problems with case?
+ *)
+fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
+ let
+ val Globals {domT, ranT, fvar, ...} = globals
+
+ val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
+
+ val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *)
+ Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
+ (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
+ (fn {prems=[a], ...} =>
+ ((rtac (G_induct OF [a]))
+ THEN_ALL_NEW (rtac accI)
+ THEN_ALL_NEW (etac R_cases)
+ THEN_ALL_NEW (asm_full_simp_tac (local_simpset_of octxt))) 1)
+
+ val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
+
+ fun mk_trsimp clause psimp =
+ let
+ val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
+ val thy = ProofContext.theory_of ctxt
+ val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+ val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
+ val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
+ fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
+ in
+ Goal.prove ctxt [] [] trsimp
+ (fn _ =>
+ rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
+ THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
+ THEN (simp_default_tac (local_simpset_of ctxt) 1)
+ THEN (etac not_acc_down 1)
+ THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (local_simpset_of ctxt))) 1)
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+ in
+ map2 mk_trsimp clauses psimps
+ end
+
+
+fun prepare_fundef config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
+ let
+ val FundefConfig {domintros, tailrec, default=default_str, ...} = config
+
+ val fvar = Free (fname, fT)
+ val domT = domain_type fT
+ val ranT = range_type fT
+
+ val default = Syntax.parse_term lthy default_str
+ |> TypeInfer.constrain fT |> Syntax.check_term lthy
+
+ val (globals, ctxt') = fix_globals domT ranT fvar lthy
+
+ val Globals { x, h, ... } = globals
+
+ val clauses = map (mk_clause_context x ctxt') abstract_qglrs
+
+ val n = length abstract_qglrs
+
+ fun build_tree (ClauseContext { ctxt, rhs, ...}) =
+ FundefCtxTree.mk_tree (fname, fT) h ctxt rhs
+
+ val trees = map build_tree clauses
+ val RCss = map find_calls trees
+
+ val ((G, GIntro_thms, G_elim, G_induct), lthy) =
+ PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
+
+ val ((f, f_defthm), lthy) =
+ PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
+
+ val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
+ val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
+
+ val ((R, RIntro_thmss, R_elim), lthy) =
+ PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
+
+ val (_, lthy) =
+ LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
+
+ val newthy = ProofContext.theory_of lthy
+ val clauses = map (transfer_clause_ctx newthy) clauses
+
+ val cert = cterm_of (ProofContext.theory_of lthy)
+
+ val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
+
+ val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
+ val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
+
+ val compat_store = store_compat_thms n compat
+
+ val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
+
+ val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
+
+ fun mk_partial_rules provedgoal =
+ let
+ val newthy = theory_of_thm provedgoal (*FIXME*)
+
+ val (graph_is_function, complete_thm) =
+ provedgoal
+ |> Conjunction.elim
+ |> apfst (Thm.forall_elim_vars 0)
+
+ val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
+
+ val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
+
+ val simple_pinduct = PROFILE "Proving partial induction rule"
+ (mk_partial_induct_rule newthy globals R complete_thm) xclauses
+
+
+ val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
+
+ val dom_intros = if domintros
+ then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
+ else NONE
+ val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
+
+ in
+ FundefResult {fs=[f], G=G, R=R, cases=complete_thm,
+ psimps=psimps, simple_pinducts=[simple_pinduct],
+ termination=total_intro, trsimps=trsimps,
+ domintros=dom_intros}
+ end
+ in
+ ((f, goalstate, mk_partial_rules), lthy)
+ end
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fundef_datatype.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,330 @@
+(* Title: HOL/Tools/Function/fundef_datatype.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+A tactic to prove completeness of datatype patterns.
+*)
+
+signature FUNDEF_DATATYPE =
+sig
+ val pat_completeness_tac: Proof.context -> int -> tactic
+ val pat_completeness: Proof.context -> Proof.method
+ val prove_completeness : theory -> term list -> term -> term list list -> term list list -> thm
+
+ val setup : theory -> theory
+
+ val add_fun : FundefCommon.fundef_config ->
+ (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
+ bool -> local_theory -> Proof.context
+ val add_fun_cmd : FundefCommon.fundef_config ->
+ (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
+ bool -> local_theory -> Proof.context
+end
+
+structure FundefDatatype : FUNDEF_DATATYPE =
+struct
+
+open FundefLib
+open FundefCommon
+
+
+fun check_pats ctxt geq =
+ let
+ fun err str = error (cat_lines ["Malformed definition:",
+ str ^ " not allowed in sequential mode.",
+ Syntax.string_of_term ctxt geq])
+ val thy = ProofContext.theory_of ctxt
+
+ fun check_constr_pattern (Bound _) = ()
+ | check_constr_pattern t =
+ let
+ val (hd, args) = strip_comb t
+ in
+ (((case Datatype.datatype_of_constr thy (fst (dest_Const hd)) of
+ SOME _ => ()
+ | NONE => err "Non-constructor pattern")
+ handle TERM ("dest_Const", _) => err "Non-constructor patterns");
+ map check_constr_pattern args;
+ ())
+ end
+
+ val (fname, qs, gs, args, rhs) = split_def ctxt geq
+
+ val _ = if not (null gs) then err "Conditional equations" else ()
+ val _ = map check_constr_pattern args
+
+ (* just count occurrences to check linearity *)
+ val _ = if fold (fold_aterms (fn Bound _ => curry (op +) 1 | _ => I)) args 0 > length qs
+ then err "Nonlinear patterns" else ()
+ in
+ ()
+ end
+
+
+fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
+fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
+
+fun inst_free var inst thm =
+ forall_elim inst (forall_intr var thm)
+
+
+fun inst_case_thm thy x P thm =
+ let
+ val [Pv, xv] = Term.add_vars (prop_of thm) []
+ in
+ cterm_instantiate [(cterm_of thy (Var xv), cterm_of thy x),
+ (cterm_of thy (Var Pv), cterm_of thy P)] thm
+ end
+
+
+fun invent_vars constr i =
+ let
+ val Ts = binder_types (fastype_of constr)
+ val j = i + length Ts
+ val is = i upto (j - 1)
+ val avs = map2 mk_argvar is Ts
+ val pvs = map2 mk_patvar is Ts
+ in
+ (avs, pvs, j)
+ end
+
+
+fun filter_pats thy cons pvars [] = []
+ | filter_pats thy cons pvars (([], thm) :: pts) = raise Match
+ | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
+ case pat of
+ Free _ => let val inst = list_comb (cons, pvars)
+ in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
+ :: (filter_pats thy cons pvars pts) end
+ | _ => if fst (strip_comb pat) = cons
+ then (pat :: pats, thm) :: (filter_pats thy cons pvars pts)
+ else filter_pats thy cons pvars pts
+
+
+fun inst_constrs_of thy (T as Type (name, _)) =
+ map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
+ (the (Datatype.get_datatype_constrs thy name))
+ | inst_constrs_of thy _ = raise Match
+
+
+fun transform_pat thy avars c_assum ([] , thm) = raise Match
+ | transform_pat thy avars c_assum (pat :: pats, thm) =
+ let
+ val (_, subps) = strip_comb pat
+ val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
+ val a_eqs = map assume eqs
+ val c_eq_pat = simplify (HOL_basic_ss addsimps a_eqs) c_assum
+ in
+ (subps @ pats, fold_rev implies_intr eqs
+ (implies_elim thm c_eq_pat))
+ end
+
+
+exception COMPLETENESS
+
+fun constr_case thy P idx (v :: vs) pats cons =
+ let
+ val (avars, pvars, newidx) = invent_vars cons idx
+ val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
+ val c_assum = assume c_hyp
+ val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
+ in
+ o_alg thy P newidx (avars @ vs) newpats
+ |> implies_intr c_hyp
+ |> fold_rev (forall_intr o cterm_of thy) avars
+ end
+ | constr_case _ _ _ _ _ _ = raise Match
+and o_alg thy P idx [] (([], Pthm) :: _) = Pthm
+ | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS
+ | o_alg thy P idx (v :: vs) pts =
+ if forall (is_Free o hd o fst) pts (* Var case *)
+ then o_alg thy P idx vs (map (fn (pv :: pats, thm) =>
+ (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts)
+ else (* Cons case *)
+ let
+ val T = fastype_of v
+ val (tname, _) = dest_Type T
+ val {exhaustion=case_thm, ...} = Datatype.the_datatype thy tname
+ val constrs = inst_constrs_of thy T
+ val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
+ in
+ inst_case_thm thy v P case_thm
+ |> fold (curry op COMP) c_cases
+ end
+ | o_alg _ _ _ _ _ = raise Match
+
+
+fun prove_completeness thy xs P qss patss =
+ let
+ fun mk_assum qs pats =
+ HOLogic.mk_Trueprop P
+ |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)
+ |> fold_rev Logic.all qs
+ |> cterm_of thy
+
+ val hyps = map2 mk_assum qss patss
+
+ fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
+
+ val assums = map2 inst_hyps hyps qss
+ in
+ o_alg thy P 2 xs (patss ~~ assums)
+ |> fold_rev implies_intr hyps
+ end
+
+
+
+fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (vs, subgf) = dest_all_all subgoal
+ val (cases, _ $ thesis) = Logic.strip_horn subgf
+ handle Bind => raise COMPLETENESS
+
+ fun pat_of assum =
+ let
+ val (qs, imp) = dest_all_all assum
+ val prems = Logic.strip_imp_prems imp
+ in
+ (qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems)
+ end
+
+ val (qss, x_pats) = split_list (map pat_of cases)
+ val xs = map fst (hd x_pats)
+ handle Empty => raise COMPLETENESS
+
+ val patss = map (map snd) x_pats
+
+ val complete_thm = prove_completeness thy xs thesis qss patss
+ |> fold_rev (forall_intr o cterm_of thy) vs
+ in
+ PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st))
+ end
+ handle COMPLETENESS => no_tac)
+
+
+fun pat_completeness ctxt = SIMPLE_METHOD' (pat_completeness_tac ctxt)
+
+val by_pat_completeness_auto =
+ Proof.global_future_terminal_proof
+ (Method.Basic (pat_completeness, Position.none),
+ SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
+
+fun termination_by method int =
+ Fundef.termination_proof NONE
+ #> Proof.global_future_terminal_proof
+ (Method.Basic (method, Position.none), NONE) int
+
+fun mk_catchall fixes arity_of =
+ let
+ fun mk_eqn ((fname, fT), _) =
+ let
+ val n = arity_of fname
+ val (argTs, rT) = chop n (binder_types fT)
+ |> apsnd (fn Ts => Ts ---> body_type fT)
+
+ val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
+ in
+ HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
+ Const ("HOL.undefined", rT))
+ |> HOLogic.mk_Trueprop
+ |> fold_rev Logic.all qs
+ end
+ in
+ map mk_eqn fixes
+ end
+
+fun add_catchall ctxt fixes spec =
+ let val fqgars = map (split_def ctxt) spec
+ val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
+ |> AList.lookup (op =) #> the
+ in
+ spec @ mk_catchall fixes arity_of
+ end
+
+fun warn_if_redundant ctxt origs tss =
+ let
+ fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
+
+ val (tss', _) = chop (length origs) tss
+ fun check (t, []) = (Output.warning (msg t); [])
+ | check (t, s) = s
+ in
+ (map check (origs ~~ tss'); tss)
+ end
+
+
+fun sequential_preproc (config as FundefConfig {sequential, ...}) ctxt fixes spec =
+ if sequential then
+ let
+ val (bnds, eqss) = split_list spec
+
+ val eqs = map the_single eqss
+
+ val feqs = eqs
+ |> tap (check_defs ctxt fixes) (* Standard checks *)
+ |> tap (map (check_pats ctxt)) (* More checks for sequential mode *)
+
+ val compleqs = add_catchall ctxt fixes feqs (* Completion *)
+
+ val spliteqs = warn_if_redundant ctxt feqs
+ (FundefSplit.split_all_equations ctxt compleqs)
+
+ fun restore_spec thms =
+ bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
+
+ val spliteqs' = flat (Library.take (length bnds, spliteqs))
+ val fnames = map (fst o fst) fixes
+ val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
+
+ fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
+ |> map (map snd)
+
+
+ val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
+
+ (* using theorem names for case name currently disabled *)
+ val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es))
+ (bnds' ~~ spliteqs)
+ |> flat
+ in
+ (flat spliteqs, restore_spec, sort, case_names)
+ end
+ else
+ FundefCommon.empty_preproc check_defs config ctxt fixes spec
+
+val setup =
+ Method.setup @{binding pat_completeness} (Scan.succeed pat_completeness)
+ "Completeness prover for datatype patterns"
+ #> Context.theory_map (FundefCommon.set_preproc sequential_preproc)
+
+
+val fun_config = FundefConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*),
+ domintros=false, tailrec=false }
+
+fun gen_fun add config fixes statements int lthy =
+ let val group = serial_string () in
+ lthy
+ |> LocalTheory.set_group group
+ |> add fixes statements config
+ |> by_pat_completeness_auto int
+ |> LocalTheory.restore
+ |> LocalTheory.set_group group
+ |> termination_by (FundefCommon.get_termination_prover lthy) int
+ end;
+
+val add_fun = gen_fun Fundef.add_fundef
+val add_fun_cmd = gen_fun Fundef.add_fundef_cmd
+
+
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+ OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
+ (fundef_parser fun_config
+ >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fundef_lib.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,176 @@
+(* Title: HOL/Tools/Function/fundef_lib.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Some fairly general functions that should probably go somewhere else...
+*)
+
+structure FundefLib = struct
+
+fun map_option f NONE = NONE
+ | map_option f (SOME x) = SOME (f x);
+
+fun fold_option f NONE y = y
+ | fold_option f (SOME x) y = f x y;
+
+fun fold_map_option f NONE y = (NONE, y)
+ | fold_map_option f (SOME x) y = apfst SOME (f x y);
+
+(* Ex: "The variable" ^ plural " is" "s are" vs *)
+fun plural sg pl [x] = sg
+ | plural sg pl _ = pl
+
+(* lambda-abstracts over an arbitrarily nested tuple
+ ==> hologic.ML? *)
+fun tupled_lambda vars t =
+ case vars of
+ (Free v) => lambda (Free v) t
+ | (Var v) => lambda (Var v) t
+ | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
+ (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
+ | _ => raise Match
+
+
+fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
+ let
+ val (n, body) = Term.dest_abs a
+ in
+ (Free (n, T), body)
+ end
+ | dest_all _ = raise Match
+
+
+(* Removes all quantifiers from a term, replacing bound variables by frees. *)
+fun dest_all_all (t as (Const ("all",_) $ _)) =
+ let
+ val (v,b) = dest_all t
+ val (vs, b') = dest_all_all b
+ in
+ (v :: vs, b')
+ end
+ | dest_all_all t = ([],t)
+
+
+(* FIXME: similar to Variable.focus *)
+fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
+ let
+ val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
+ val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
+
+ val (n'', body) = Term.dest_abs (n', T, b)
+ val _ = (n' = n'') orelse error "dest_all_ctx"
+ (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
+
+ val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
+ in
+ (ctx'', (n', T) :: vs, bd)
+ end
+ | dest_all_all_ctx ctx t =
+ (ctx, [], t)
+
+
+fun map3 _ [] [] [] = []
+ | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
+ | map3 _ _ _ _ = raise Library.UnequalLengths;
+
+fun map4 _ [] [] [] [] = []
+ | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
+ | map4 _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map6 _ [] [] [] [] [] [] = []
+ | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
+ | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map7 _ [] [] [] [] [] [] [] = []
+ | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
+ | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+
+
+(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
+(* ==> library *)
+fun unordered_pairs [] = []
+ | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
+
+
+(* Replaces Frees by name. Works with loose Bounds. *)
+fun replace_frees assoc =
+ map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
+ | t => t)
+
+
+fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
+ | rename_bound n _ = raise Match
+
+fun mk_forall_rename (n, v) =
+ rename_bound n o Logic.all v
+
+fun forall_intr_rename (n, cv) thm =
+ let
+ val allthm = forall_intr cv thm
+ val (_ $ abs) = prop_of allthm
+ in
+ Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
+ end
+
+
+(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
+fun frees_in_term ctxt t =
+ Term.add_frees t []
+ |> filter_out (Variable.is_fixed ctxt o fst)
+ |> rev
+
+
+datatype proof_attempt = Solved of thm | Stuck of thm | Fail
+
+fun try_proof cgoal tac =
+ case SINGLE tac (Goal.init cgoal) of
+ NONE => Fail
+ | SOME st => if Thm.no_prems st then Solved (Goal.finish st) else Stuck st
+
+
+fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) =
+ if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
+ | dest_binop_list _ t = [ t ]
+
+
+(* separate two parts in a +-expression:
+ "a + b + c + d + e" --> "(a + b + d) + (c + e)"
+
+ Here, + can be any binary operation that is AC.
+
+ cn - The name of the binop-constructor (e.g. @{const_name Un})
+ ac - the AC rewrite rules for cn
+ is - the list of indices of the expressions that should become the first part
+ (e.g. [0,1,3] in the above example)
+*)
+
+fun regroup_conv neu cn ac is ct =
+ let
+ val mk = HOLogic.mk_binop cn
+ val t = term_of ct
+ val xs = dest_binop_list cn t
+ val js = 0 upto (length xs) - 1 \\ is
+ val ty = fastype_of t
+ val thy = theory_of_cterm ct
+ in
+ Goal.prove_internal []
+ (cterm_of thy
+ (Logic.mk_equals (t,
+ if is = []
+ then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
+ else if js = []
+ then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
+ else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
+ (K (rewrite_goals_tac ac
+ THEN rtac Drule.reflexive_thm 1))
+ end
+
+(* instance for unions *)
+fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Un}
+ (map (fn t => t RS eq_reflection) (@{thms "Un_ac"} @
+ @{thms "Un_empty_right"} @
+ @{thms "Un_empty_left"})) t
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/induction_scheme.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,405 @@
+(* Title: HOL/Tools/Function/induction_scheme.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A method to prove induction schemes.
+*)
+
+signature INDUCTION_SCHEME =
+sig
+ val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic)
+ -> Proof.context -> thm list -> tactic
+ val induct_scheme_tac : Proof.context -> thm list -> tactic
+ val setup : theory -> theory
+end
+
+
+structure InductionScheme : INDUCTION_SCHEME =
+struct
+
+open FundefLib
+
+
+type rec_call_info = int * (string * typ) list * term list * term list
+
+datatype scheme_case =
+ SchemeCase of
+ {
+ bidx : int,
+ qs: (string * typ) list,
+ oqnames: string list,
+ gs: term list,
+ lhs: term list,
+ rs: rec_call_info list
+ }
+
+datatype scheme_branch =
+ SchemeBranch of
+ {
+ P : term,
+ xs: (string * typ) list,
+ ws: (string * typ) list,
+ Cs: term list
+ }
+
+datatype ind_scheme =
+ IndScheme of
+ {
+ T: typ, (* sum of products *)
+ branches: scheme_branch list,
+ cases: scheme_case list
+ }
+
+val ind_atomize = MetaSimplifier.rewrite true @{thms induct_atomize}
+val ind_rulify = MetaSimplifier.rewrite true @{thms induct_rulify}
+
+fun meta thm = thm RS eq_reflection
+
+val sum_prod_conv = MetaSimplifier.rewrite true
+ (map meta (@{thm split_conv} :: @{thms sum.cases}))
+
+fun term_conv thy cv t =
+ cv (cterm_of thy t)
+ |> prop_of |> Logic.dest_equals |> snd
+
+fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T))
+
+fun dest_hhf ctxt t =
+ let
+ val (ctxt', vars, imp) = dest_all_all_ctx ctxt t
+ in
+ (ctxt', vars, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp)
+ end
+
+
+fun mk_scheme' ctxt cases concl =
+ let
+ fun mk_branch concl =
+ let
+ val (ctxt', ws, Cs, _ $ Pxs) = dest_hhf ctxt concl
+ val (P, xs) = strip_comb Pxs
+ in
+ SchemeBranch { P=P, xs=map dest_Free xs, ws=ws, Cs=Cs }
+ end
+
+ val (branches, cases') = (* correction *)
+ case Logic.dest_conjunction_list concl of
+ [conc] =>
+ let
+ val _ $ Pxs = Logic.strip_assums_concl conc
+ val (P, _) = strip_comb Pxs
+ val (cases', conds) = take_prefix (Term.exists_subterm (curry op aconv P)) cases
+ val concl' = fold_rev (curry Logic.mk_implies) conds conc
+ in
+ ([mk_branch concl'], cases')
+ end
+ | concls => (map mk_branch concls, cases)
+
+ fun mk_case premise =
+ let
+ val (ctxt', qs, prems, _ $ Plhs) = dest_hhf ctxt premise
+ val (P, lhs) = strip_comb Plhs
+
+ fun bidx Q = find_index (fn SchemeBranch {P=P',...} => Q aconv P') branches
+
+ fun mk_rcinfo pr =
+ let
+ val (ctxt'', Gvs, Gas, _ $ Phyp) = dest_hhf ctxt' pr
+ val (P', rcs) = strip_comb Phyp
+ in
+ (bidx P', Gvs, Gas, rcs)
+ end
+
+ fun is_pred v = exists (fn SchemeBranch {P,...} => v aconv P) branches
+
+ val (gs, rcprs) =
+ take_prefix (not o Term.exists_subterm is_pred) prems
+ in
+ SchemeCase {bidx=bidx P, qs=qs, oqnames=map fst qs(*FIXME*), gs=gs, lhs=lhs, rs=map mk_rcinfo rcprs}
+ end
+
+ fun PT_of (SchemeBranch { xs, ...}) =
+ foldr1 HOLogic.mk_prodT (map snd xs)
+
+ val ST = BalancedTree.make (uncurry SumTree.mk_sumT) (map PT_of branches)
+ in
+ IndScheme {T=ST, cases=map mk_case cases', branches=branches }
+ end
+
+
+
+fun mk_completeness ctxt (IndScheme {cases, branches, ...}) bidx =
+ let
+ val SchemeBranch { xs, ws, Cs, ... } = nth branches bidx
+ val relevant_cases = filter (fn SchemeCase {bidx=bidx', ...} => bidx' = bidx) cases
+
+ val allqnames = fold (fn SchemeCase {qs, ...} => fold (insert (op =) o Free) qs) relevant_cases []
+ val (Pbool :: xs') = map Free (Variable.variant_frees ctxt allqnames (("P", HOLogic.boolT) :: xs))
+ val Cs' = map (Pattern.rewrite_term (ProofContext.theory_of ctxt) (filter_out (op aconv) (map Free xs ~~ xs')) []) Cs
+
+ fun mk_case (SchemeCase {qs, oqnames, gs, lhs, ...}) =
+ HOLogic.mk_Trueprop Pbool
+ |> fold_rev (fn x_l => curry Logic.mk_implies (HOLogic.mk_Trueprop(HOLogic.mk_eq x_l)))
+ (xs' ~~ lhs)
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
+ in
+ HOLogic.mk_Trueprop Pbool
+ |> fold_rev (curry Logic.mk_implies o mk_case) relevant_cases
+ |> fold_rev (curry Logic.mk_implies) Cs'
+ |> fold_rev (Logic.all o Free) ws
+ |> fold_rev mk_forall_rename (map fst xs ~~ xs')
+ |> mk_forall_rename ("P", Pbool)
+ end
+
+fun mk_wf ctxt R (IndScheme {T, ...}) =
+ HOLogic.Trueprop $ (Const (@{const_name "wf"}, mk_relT T --> HOLogic.boolT) $ R)
+
+fun mk_ineqs R (IndScheme {T, cases, branches}) =
+ let
+ fun inject i ts =
+ SumTree.mk_inj T (length branches) (i + 1) (foldr1 HOLogic.mk_prod ts)
+
+ val thesis = Free ("thesis", HOLogic.boolT) (* FIXME *)
+
+ fun mk_pres bdx args =
+ let
+ val SchemeBranch { xs, ws, Cs, ... } = nth branches bdx
+ fun replace (x, v) t = betapply (lambda (Free x) t, v)
+ val Cs' = map (fold replace (xs ~~ args)) Cs
+ val cse =
+ HOLogic.mk_Trueprop thesis
+ |> fold_rev (curry Logic.mk_implies) Cs'
+ |> fold_rev (Logic.all o Free) ws
+ in
+ Logic.mk_implies (cse, HOLogic.mk_Trueprop thesis)
+ end
+
+ fun f (SchemeCase {bidx, qs, oqnames, gs, lhs, rs, ...}) =
+ let
+ fun g (bidx', Gvs, Gas, rcarg) =
+ let val export =
+ fold_rev (curry Logic.mk_implies) Gas
+ #> fold_rev (curry Logic.mk_implies) gs
+ #> fold_rev (Logic.all o Free) Gvs
+ #> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
+ in
+ (HOLogic.mk_mem (HOLogic.mk_prod (inject bidx' rcarg, inject bidx lhs), R)
+ |> HOLogic.mk_Trueprop
+ |> export,
+ mk_pres bidx' rcarg
+ |> export
+ |> Logic.all thesis)
+ end
+ in
+ map g rs
+ end
+ in
+ map f cases
+ end
+
+
+fun mk_hol_imp a b = HOLogic.imp $ a $ b
+
+fun mk_ind_goal thy branches =
+ let
+ fun brnch (SchemeBranch { P, xs, ws, Cs, ... }) =
+ HOLogic.mk_Trueprop (list_comb (P, map Free xs))
+ |> fold_rev (curry Logic.mk_implies) Cs
+ |> fold_rev (Logic.all o Free) ws
+ |> term_conv thy ind_atomize
+ |> ObjectLogic.drop_judgment thy
+ |> tupled_lambda (foldr1 HOLogic.mk_prod (map Free xs))
+ in
+ SumTree.mk_sumcases HOLogic.boolT (map brnch branches)
+ end
+
+
+fun mk_induct_rule ctxt R x complete_thms wf_thm ineqss (IndScheme {T, cases=scases, branches}) =
+ let
+ val n = length branches
+
+ val scases_idx = map_index I scases
+
+ fun inject i ts =
+ SumTree.mk_inj T n (i + 1) (foldr1 HOLogic.mk_prod ts)
+ val P_of = nth (map (fn (SchemeBranch { P, ... }) => P) branches)
+
+ val thy = ProofContext.theory_of ctxt
+ val cert = cterm_of thy
+
+ val P_comp = mk_ind_goal thy branches
+
+ (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
+ val ihyp = Term.all T $ Abs ("z", T,
+ Logic.mk_implies
+ (HOLogic.mk_Trueprop (
+ Const ("op :", HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT)
+ $ (HOLogic.pair_const T T $ Bound 0 $ x)
+ $ R),
+ HOLogic.mk_Trueprop (P_comp $ Bound 0)))
+ |> cert
+
+ val aihyp = assume ihyp
+
+ (* Rule for case splitting along the sum types *)
+ val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
+ val pats = map_index (uncurry inject) xss
+ val sum_split_rule = FundefDatatype.prove_completeness thy [x] (P_comp $ x) xss (map single pats)
+
+ fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
+ let
+ val fxs = map Free xs
+ val branch_hyp = assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat))))
+
+ val C_hyps = map (cert #> assume) Cs
+
+ val (relevant_cases, ineqss') = filter (fn ((_, SchemeCase {bidx=bidx', ...}), _) => bidx' = bidx) (scases_idx ~~ ineqss)
+ |> split_list
+
+ fun prove_case (cidx, SchemeCase {qs, oqnames, gs, lhs, rs, ...}) ineq_press =
+ let
+ val case_hyps = map (assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs)
+
+ val cqs = map (cert o Free) qs
+ val ags = map (assume o cert) gs
+
+ val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps)
+ val sih = full_simplify replace_x_ss aihyp
+
+ fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) =
+ let
+ val cGas = map (assume o cert) Gas
+ val cGvs = map (cert o Free) Gvs
+ val import = fold forall_elim (cqs @ cGvs)
+ #> fold Thm.elim_implies (ags @ cGas)
+ val ipres = pres
+ |> forall_elim (cert (list_comb (P_of idx, rcargs)))
+ |> import
+ in
+ sih |> forall_elim (cert (inject idx rcargs))
+ |> Thm.elim_implies (import ineq) (* Psum rcargs *)
+ |> Conv.fconv_rule sum_prod_conv
+ |> Conv.fconv_rule ind_rulify
+ |> (fn th => th COMP ipres) (* P rs *)
+ |> fold_rev (implies_intr o cprop_of) cGas
+ |> fold_rev forall_intr cGvs
+ end
+
+ val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *)
+
+ val step = HOLogic.mk_Trueprop (list_comb (P, lhs))
+ |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev (Logic.all o Free) qs
+ |> cert
+
+ val Plhs_to_Pxs_conv =
+ foldl1 (uncurry Conv.combination_conv)
+ (Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps)
+
+ val res = assume step
+ |> fold forall_elim cqs
+ |> fold Thm.elim_implies ags
+ |> fold Thm.elim_implies P_recs (* P lhs *)
+ |> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *)
+ |> fold_rev (implies_intr o cprop_of) (ags @ case_hyps)
+ |> fold_rev forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
+ in
+ (res, (cidx, step))
+ end
+
+ val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss')
+
+ val bstep = complete_thm
+ |> forall_elim (cert (list_comb (P, fxs)))
+ |> fold (forall_elim o cert) (fxs @ map Free ws)
+ |> fold Thm.elim_implies C_hyps (* FIXME: optimization using rotate_prems *)
+ |> fold Thm.elim_implies cases (* P xs *)
+ |> fold_rev (implies_intr o cprop_of) C_hyps
+ |> fold_rev (forall_intr o cert o Free) ws
+
+ val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x))
+ |> Goal.init
+ |> (MetaSimplifier.rewrite_goals_tac (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.cases}))
+ THEN CONVERSION ind_rulify 1)
+ |> Seq.hd
+ |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
+ |> Goal.finish
+ |> implies_intr (cprop_of branch_hyp)
+ |> fold_rev (forall_intr o cert) fxs
+ in
+ (Pxs, steps)
+ end
+
+ val (branches, steps) = split_list (map_index prove_branch (branches ~~ (complete_thms ~~ pats)))
+ |> apsnd flat
+
+ val istep = sum_split_rule
+ |> fold (fn b => fn th => Drule.compose_single (b, 1, th)) branches
+ |> implies_intr ihyp
+ |> forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *)
+
+ val induct_rule =
+ @{thm "wf_induct_rule"}
+ |> (curry op COMP) wf_thm
+ |> (curry op COMP) istep
+
+ val steps_sorted = map snd (sort (int_ord o pairself fst) steps)
+ in
+ (steps_sorted, induct_rule)
+ end
+
+
+fun mk_ind_tac comp_tac pres_tac term_tac ctxt facts = (ALLGOALS (Method.insert_tac facts)) THEN HEADGOAL
+(SUBGOAL (fn (t, i) =>
+ let
+ val (ctxt', _, cases, concl) = dest_hhf ctxt t
+ val scheme as IndScheme {T=ST, branches, ...} = mk_scheme' ctxt' cases concl
+(* val _ = Output.tracing (makestring scheme)*)
+ val ([Rn,xn], ctxt'') = Variable.variant_fixes ["R","x"] ctxt'
+ val R = Free (Rn, mk_relT ST)
+ val x = Free (xn, ST)
+ val cert = cterm_of (ProofContext.theory_of ctxt)
+
+ val ineqss = mk_ineqs R scheme
+ |> map (map (pairself (assume o cert)))
+ val complete = map (mk_completeness ctxt scheme #> cert #> assume) (0 upto (length branches - 1))
+ val wf_thm = mk_wf ctxt R scheme |> cert |> assume
+
+ val (descent, pres) = split_list (flat ineqss)
+ val newgoals = complete @ pres @ wf_thm :: descent
+
+ val (steps, indthm) = mk_induct_rule ctxt'' R x complete wf_thm ineqss scheme
+
+ fun project (i, SchemeBranch {xs, ...}) =
+ let
+ val inst = cert (SumTree.mk_inj ST (length branches) (i + 1) (foldr1 HOLogic.mk_prod (map Free xs)))
+ in
+ indthm |> Drule.instantiate' [] [SOME inst]
+ |> simplify SumTree.sumcase_split_ss
+ |> Conv.fconv_rule ind_rulify
+(* |> (fn thm => (Output.tracing (makestring thm); thm))*)
+ end
+
+ val res = Conjunction.intr_balanced (map_index project branches)
+ |> fold_rev implies_intr (map cprop_of newgoals @ steps)
+ |> (fn thm => Thm.generalize ([], [Rn]) (Thm.maxidx_of thm + 1) thm)
+
+ val nbranches = length branches
+ val npres = length pres
+ in
+ Thm.compose_no_flatten false (res, length newgoals) i
+ THEN term_tac (i + nbranches + npres)
+ THEN (EVERY (map (TRY o pres_tac) ((i + nbranches + npres - 1) downto (i + nbranches))))
+ THEN (EVERY (map (TRY o comp_tac) ((i + nbranches - 1) downto i)))
+ end))
+
+
+fun induct_scheme_tac ctxt =
+ mk_ind_tac (K all_tac) (assume_tac APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt;
+
+val setup =
+ Method.setup @{binding induct_scheme} (Scan.succeed (RAW_METHOD o induct_scheme_tac))
+ "proves an induction principle"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/inductive_wrap.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,67 @@
+(* Title: HOL/Tools/Function/inductive_wrap.ML
+ Author: Alexander Krauss, TU Muenchen
+
+
+A wrapper around the inductive package, restoring the quantifiers in
+the introduction and elimination rules.
+*)
+
+signature FUNDEF_INDUCTIVE_WRAP =
+sig
+ val inductive_def : term list
+ -> ((bstring * typ) * mixfix) * local_theory
+ -> thm list * (term * thm * thm * local_theory)
+end
+
+structure FundefInductiveWrap: FUNDEF_INDUCTIVE_WRAP =
+struct
+
+open FundefLib
+
+fun requantify ctxt lfix orig_def thm =
+ let
+ val (qs, t) = dest_all_all orig_def
+ val thy = theory_of_thm thm
+ val frees = frees_in_term ctxt t
+ |> remove (op =) lfix
+ val vars = Term.add_vars (prop_of thm) [] |> rev
+
+ val varmap = frees ~~ vars
+ in
+ fold_rev (fn Free (n, T) =>
+ forall_intr_rename (n, cterm_of thy (Var (the_default (("",0), T) (AList.lookup (op =) varmap (n, T))))))
+ qs
+ thm
+ end
+
+
+
+fun inductive_def defs (((R, T), mixfix), lthy) =
+ let
+ val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) =
+ Inductive.add_inductive_i
+ {quiet_mode = false,
+ verbose = ! Toplevel.debug,
+ kind = Thm.internalK,
+ alt_name = Binding.empty,
+ coind = false,
+ no_elim = false,
+ no_ind = false,
+ skip_mono = true,
+ fork_mono = false}
+ [((Binding.name R, T), NoSyn)] (* the relation *)
+ [] (* no parameters *)
+ (map (fn t => (Attrib.empty_binding, t)) defs) (* the intros *)
+ [] (* no special monos *)
+ lthy
+
+ val intrs = map2 (requantify lthy (R, T)) defs intrs_gen
+
+ val elim = elim_gen
+ |> forall_intr_vars (* FIXME... *)
+
+ in
+ (intrs, (Rdef, elim, induct, lthy))
+ end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/lexicographic_order.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,230 @@
+(* Title: HOL/Tools/Function/lexicographic_order.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Method for termination proofs with lexicographic orderings.
+*)
+
+signature LEXICOGRAPHIC_ORDER =
+sig
+ val lex_order_tac : Proof.context -> tactic -> tactic
+ val lexicographic_order_tac : Proof.context -> tactic
+ val lexicographic_order : Proof.context -> Proof.method
+
+ val setup: theory -> theory
+end
+
+structure LexicographicOrder : LEXICOGRAPHIC_ORDER =
+struct
+
+open FundefLib
+
+(** General stuff **)
+
+fun mk_measures domT mfuns =
+ let
+ val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))
+ val mlexT = (domT --> HOLogic.natT) --> relT --> relT
+ fun mk_ms [] = Const (@{const_name Set.empty}, relT)
+ | mk_ms (f::fs) =
+ Const (@{const_name "mlex_prod"}, mlexT) $ f $ mk_ms fs
+ in
+ mk_ms mfuns
+ end
+
+fun del_index n [] = []
+ | del_index n (x :: xs) =
+ if n > 0 then x :: del_index (n - 1) xs else xs
+
+fun transpose ([]::_) = []
+ | transpose xss = map hd xss :: transpose (map tl xss)
+
+(** Matrix cell datatype **)
+
+datatype cell = Less of thm| LessEq of (thm * thm) | None of (thm * thm) | False of thm;
+
+fun is_Less (Less _) = true
+ | is_Less _ = false
+
+fun is_LessEq (LessEq _) = true
+ | is_LessEq _ = false
+
+fun pr_cell (Less _ ) = " < "
+ | pr_cell (LessEq _) = " <="
+ | pr_cell (None _) = " ? "
+ | pr_cell (False _) = " F "
+
+
+(** Proof attempts to build the matrix **)
+
+fun dest_term (t : term) =
+ let
+ val (vars, prop) = FundefLib.dest_all_all t
+ val (prems, concl) = Logic.strip_horn prop
+ val (lhs, rhs) = concl
+ |> HOLogic.dest_Trueprop
+ |> HOLogic.dest_mem |> fst
+ |> HOLogic.dest_prod
+ in
+ (vars, prems, lhs, rhs)
+ end
+
+fun mk_goal (vars, prems, lhs, rhs) rel =
+ let
+ val concl = HOLogic.mk_binrel rel (lhs, rhs) |> HOLogic.mk_Trueprop
+ in
+ fold_rev Logic.all vars (Logic.list_implies (prems, concl))
+ end
+
+fun prove thy solve_tac t =
+ cterm_of thy t |> Goal.init
+ |> SINGLE solve_tac |> the
+
+fun mk_cell (thy : theory) solve_tac (vars, prems, lhs, rhs) mfun =
+ let
+ val goals = cterm_of thy o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs)
+ in
+ case try_proof (goals @{const_name HOL.less}) solve_tac of
+ Solved thm => Less thm
+ | Stuck thm =>
+ (case try_proof (goals @{const_name HOL.less_eq}) solve_tac of
+ Solved thm2 => LessEq (thm2, thm)
+ | Stuck thm2 =>
+ if prems_of thm2 = [HOLogic.Trueprop $ HOLogic.false_const] then False thm2
+ else None (thm2, thm)
+ | _ => raise Match) (* FIXME *)
+ | _ => raise Match
+ end
+
+
+(** Search algorithms **)
+
+fun check_col ls = forall (fn c => is_Less c orelse is_LessEq c) ls andalso not (forall (is_LessEq) ls)
+
+fun transform_table table col = table |> filter_out (fn x => is_Less (nth x col)) |> map (del_index col)
+
+fun transform_order col order = map (fn x => if x >= col then x + 1 else x) order
+
+(* simple depth-first search algorithm for the table *)
+fun search_table table =
+ case table of
+ [] => SOME []
+ | _ =>
+ let
+ val col = find_index (check_col) (transpose table)
+ in case col of
+ ~1 => NONE
+ | _ =>
+ let
+ val order_opt = (table, col) |-> transform_table |> search_table
+ in case order_opt of
+ NONE => NONE
+ | SOME order =>SOME (col :: transform_order col order)
+ end
+ end
+
+(** Proof Reconstruction **)
+
+(* prove row :: cell list -> tactic *)
+fun prove_row (Less less_thm :: _) =
+ (rtac @{thm "mlex_less"} 1)
+ THEN PRIMITIVE (Thm.elim_implies less_thm)
+ | prove_row (LessEq (lesseq_thm, _) :: tail) =
+ (rtac @{thm "mlex_leq"} 1)
+ THEN PRIMITIVE (Thm.elim_implies lesseq_thm)
+ THEN prove_row tail
+ | prove_row _ = sys_error "lexicographic_order"
+
+
+(** Error reporting **)
+
+fun pr_table table = writeln (cat_lines (map (fn r => concat (map pr_cell r)) table))
+
+fun pr_goals ctxt st =
+ Display.pretty_goals_aux (Syntax.pp ctxt) Markup.none (true, false) (Thm.nprems_of st) st
+ |> Pretty.chunks
+ |> Pretty.string_of
+
+fun row_index i = chr (i + 97)
+fun col_index j = string_of_int (j + 1)
+
+fun pr_unprovable_cell _ ((i,j), Less _) = ""
+ | pr_unprovable_cell ctxt ((i,j), LessEq (_, st)) =
+ "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
+ | pr_unprovable_cell ctxt ((i,j), None (st_leq, st_less)) =
+ "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st_less
+ ^ "\n(" ^ row_index i ^ ", " ^ col_index j ^ ", <=):\n" ^ pr_goals ctxt st_leq
+ | pr_unprovable_cell ctxt ((i,j), False st) =
+ "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
+
+fun pr_unprovable_subgoals ctxt table =
+ table
+ |> map_index (fn (i,cs) => map_index (fn (j,x) => ((i,j), x)) cs)
+ |> flat
+ |> map (pr_unprovable_cell ctxt)
+
+fun no_order_msg ctxt table tl measure_funs =
+ let
+ val prterm = Syntax.string_of_term ctxt
+ fun pr_fun t i = string_of_int i ^ ") " ^ prterm t
+
+ fun pr_goal t i =
+ let
+ val (_, _, lhs, rhs) = dest_term t
+ in (* also show prems? *)
+ i ^ ") " ^ prterm rhs ^ " ~> " ^ prterm lhs
+ end
+
+ val gc = map (fn i => chr (i + 96)) (1 upto length table)
+ val mc = 1 upto length measure_funs
+ val tstr = "Result matrix:" :: (" " ^ concat (map (enclose " " " " o string_of_int) mc))
+ :: map2 (fn r => fn i => i ^ ": " ^ concat (map pr_cell r)) table gc
+ val gstr = "Calls:" :: map2 (prefix " " oo pr_goal) tl gc
+ val mstr = "Measures:" :: map2 (prefix " " oo pr_fun) measure_funs mc
+ val ustr = "Unfinished subgoals:" :: pr_unprovable_subgoals ctxt table
+ in
+ cat_lines (ustr @ gstr @ mstr @ tstr @ ["", "Could not find lexicographic termination order."])
+ end
+
+(** The Main Function **)
+
+fun lex_order_tac ctxt solve_tac (st: thm) =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val ((trueprop $ (wf $ rel)) :: tl) = prems_of st
+
+ val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel))
+
+ val measure_funs = MeasureFunctions.get_measure_functions ctxt domT (* 1: generate measures *)
+
+ (* 2: create table *)
+ val table = map (fn t => map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl
+
+ val order = the (search_table table) (* 3: search table *)
+ handle Option => error (no_order_msg ctxt table tl measure_funs)
+
+ val clean_table = map (fn x => map (nth x) order) table
+
+ val relation = mk_measures domT (map (nth measure_funs) order)
+ val _ = writeln ("Found termination order: " ^ quote (Syntax.string_of_term ctxt relation))
+
+ in (* 4: proof reconstruction *)
+ st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)])
+ THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
+ THEN (rtac @{thm "wf_empty"} 1)
+ THEN EVERY (map prove_row clean_table))
+ end
+
+fun lexicographic_order_tac ctxt =
+ TRY (FundefCommon.apply_termination_rule ctxt 1)
+ THEN lex_order_tac ctxt (auto_tac (local_clasimpset_of ctxt addsimps2 FundefCommon.TerminationSimps.get ctxt))
+
+val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac
+
+val setup =
+ Method.setup @{binding lexicographic_order}
+ (Method.sections clasimp_modifiers >> (K lexicographic_order))
+ "termination prover for lexicographic orderings"
+ #> Context.theory_map (FundefCommon.set_termination_prover lexicographic_order)
+
+end;
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/measure_functions.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,58 @@
+(* Title: HOL/Tools/Function/measure_functions.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Measure functions, generated heuristically
+*)
+
+signature MEASURE_FUNCTIONS =
+sig
+
+ val get_measure_functions : Proof.context -> typ -> term list
+ val setup : theory -> theory
+
+end
+
+structure MeasureFunctions : MEASURE_FUNCTIONS =
+struct
+
+(** User-declared size functions **)
+structure MeasureHeuristicRules = NamedThmsFun(
+ val name = "measure_function"
+ val description = "Rules that guide the heuristic generation of measure functions"
+);
+
+fun mk_is_measures t = Const (@{const_name "is_measure"}, fastype_of t --> HOLogic.boolT) $ t
+
+fun find_measures ctxt T =
+ DEPTH_SOLVE (resolve_tac (MeasureHeuristicRules.get ctxt) 1)
+ (HOLogic.mk_Trueprop (mk_is_measures (Var (("f",0), T --> HOLogic.natT)))
+ |> cterm_of (ProofContext.theory_of ctxt) |> Goal.init)
+ |> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
+ |> Seq.list_of
+
+
+(** Generating Measure Functions **)
+
+fun constant_0 T = Abs ("x", T, HOLogic.zero)
+fun constant_1 T = Abs ("x", T, HOLogic.Suc_zero)
+
+fun mk_funorder_funs (Type ("+", [fT, sT])) =
+ map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT m (constant_0 sT)) (mk_funorder_funs fT)
+ @ map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT (constant_0 fT) m) (mk_funorder_funs sT)
+ | mk_funorder_funs T = [ constant_1 T ]
+
+fun mk_ext_base_funs ctxt (Type("+", [fT, sT])) =
+ map_product (SumTree.mk_sumcase fT sT HOLogic.natT)
+ (mk_ext_base_funs ctxt fT) (mk_ext_base_funs ctxt sT)
+ | mk_ext_base_funs ctxt T = find_measures ctxt T
+
+fun mk_all_measure_funs ctxt (T as Type ("+", _)) =
+ mk_ext_base_funs ctxt T @ mk_funorder_funs T
+ | mk_all_measure_funs ctxt T = find_measures ctxt T
+
+val get_measure_functions = mk_all_measure_funs
+
+val setup = MeasureHeuristicRules.setup
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/mutual.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,314 @@
+(* Title: HOL/Tools/Function/mutual.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Tools for mutual recursive definitions.
+*)
+
+signature FUNDEF_MUTUAL =
+sig
+
+ val prepare_fundef_mutual : FundefCommon.fundef_config
+ -> string (* defname *)
+ -> ((string * typ) * mixfix) list
+ -> term list
+ -> local_theory
+ -> ((thm (* goalstate *)
+ * (thm -> FundefCommon.fundef_result) (* proof continuation *)
+ ) * local_theory)
+
+end
+
+
+structure FundefMutual: FUNDEF_MUTUAL =
+struct
+
+open FundefLib
+open FundefCommon
+
+
+
+
+type qgar = string * (string * typ) list * term list * term list * term
+
+fun name_of_fqgar ((f, _, _, _, _): qgar) = f
+
+datatype mutual_part =
+ MutualPart of
+ {
+ i : int,
+ i' : int,
+ fvar : string * typ,
+ cargTs: typ list,
+ f_def: term,
+
+ f: term option,
+ f_defthm : thm option
+ }
+
+
+datatype mutual_info =
+ Mutual of
+ {
+ n : int,
+ n' : int,
+ fsum_var : string * typ,
+
+ ST: typ,
+ RST: typ,
+
+ parts: mutual_part list,
+ fqgars: qgar list,
+ qglrs: ((string * typ) list * term list * term * term) list,
+
+ fsum : term option
+ }
+
+fun mutual_induct_Pnames n =
+ if n < 5 then fst (chop n ["P","Q","R","S"])
+ else map (fn i => "P" ^ string_of_int i) (1 upto n)
+
+fun get_part fname =
+ the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
+
+(* FIXME *)
+fun mk_prod_abs e (t1, t2) =
+ let
+ val bTs = rev (map snd e)
+ val T1 = fastype_of1 (bTs, t1)
+ val T2 = fastype_of1 (bTs, t2)
+ in
+ HOLogic.pair_const T1 T2 $ t1 $ t2
+ end;
+
+
+fun analyze_eqs ctxt defname fs eqs =
+ let
+ val num = length fs
+ val fnames = map fst fs
+ val fqgars = map (split_def ctxt) eqs
+ val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
+ |> AList.lookup (op =) #> the
+
+ fun curried_types (fname, fT) =
+ let
+ val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
+ in
+ (caTs, uaTs ---> body_type fT)
+ end
+
+ val (caTss, resultTs) = split_list (map curried_types fs)
+ val argTs = map (foldr1 HOLogic.mk_prodT) caTss
+
+ val dresultTs = distinct (Type.eq_type Vartab.empty) resultTs
+ val n' = length dresultTs
+
+ val RST = BalancedTree.make (uncurry SumTree.mk_sumT) dresultTs
+ val ST = BalancedTree.make (uncurry SumTree.mk_sumT) argTs
+
+ val fsum_type = ST --> RST
+
+ val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
+ val fsum_var = (fsum_var_name, fsum_type)
+
+ fun define (fvar as (n, T)) caTs resultT i =
+ let
+ val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
+ val i' = find_index (fn Ta => Type.eq_type Vartab.empty (Ta, resultT)) dresultTs + 1
+
+ val f_exp = SumTree.mk_proj RST n' i' (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
+ val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
+
+ val rew = (n, fold_rev lambda vars f_exp)
+ in
+ (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
+ end
+
+ val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
+
+ fun convert_eqs (f, qs, gs, args, rhs) =
+ let
+ val MutualPart {i, i', ...} = get_part f parts
+ in
+ (qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
+ SumTree.mk_inj RST n' i' (replace_frees rews rhs)
+ |> Envir.beta_norm)
+ end
+
+ val qglrs = map convert_eqs fqgars
+ in
+ Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST,
+ parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
+ end
+
+
+
+
+fun define_projections fixes mutual fsum lthy =
+ let
+ fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
+ let
+ val ((f, (_, f_defthm)), lthy') =
+ LocalTheory.define Thm.internalK ((Binding.name fname, mixfix),
+ ((Binding.name (fname ^ "_def"), []), Term.subst_bound (fsum, f_def)))
+ lthy
+ in
+ (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
+ f=SOME f, f_defthm=SOME f_defthm },
+ lthy')
+ end
+
+ val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
+ val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
+ in
+ (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
+ fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
+ lthy')
+ end
+
+
+fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
+ let
+ val thy = ProofContext.theory_of ctxt
+
+ val oqnames = map fst pre_qs
+ val (qs, ctxt') = Variable.variant_fixes oqnames ctxt
+ |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
+
+ fun inst t = subst_bounds (rev qs, t)
+ val gs = map inst pre_gs
+ val args = map inst pre_args
+ val rhs = inst pre_rhs
+
+ val cqs = map (cterm_of thy) qs
+ val ags = map (assume o cterm_of thy) gs
+
+ val import = fold forall_elim cqs
+ #> fold Thm.elim_implies ags
+
+ val export = fold_rev (implies_intr o cprop_of) ags
+ #> fold_rev forall_intr_rename (oqnames ~~ cqs)
+ in
+ F ctxt (f, qs, gs, args, rhs) import export
+ end
+
+fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs) import (export : thm -> thm) sum_psimp_eq =
+ let
+ val (MutualPart {f=SOME f, f_defthm=SOME f_def, ...}) = get_part fname parts
+
+ val psimp = import sum_psimp_eq
+ val (simp, restore_cond) = case cprems_of psimp of
+ [] => (psimp, I)
+ | [cond] => (implies_elim psimp (assume cond), implies_intr cond)
+ | _ => sys_error "Too many conditions"
+ in
+ Goal.prove ctxt [] []
+ (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
+ (fn _ => (LocalDefs.unfold_tac ctxt all_orig_fdefs)
+ THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
+ THEN (simp_tac (local_simpset_of ctxt addsimps SumTree.proj_in_rules)) 1)
+ |> restore_cond
+ |> export
+ end
+
+
+(* FIXME HACK *)
+fun mk_applied_form ctxt caTs thm =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
+ in
+ fold (fn x => fn thm => combination thm (reflexive x)) xs thm
+ |> Conv.fconv_rule (Thm.beta_conversion true)
+ |> fold_rev forall_intr xs
+ |> Thm.forall_elim_vars 0
+ end
+
+
+fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, RST, parts, ...}) =
+ let
+ val cert = cterm_of (ProofContext.theory_of lthy)
+ val newPs = map2 (fn Pname => fn MutualPart {cargTs, ...} =>
+ Free (Pname, cargTs ---> HOLogic.boolT))
+ (mutual_induct_Pnames (length parts))
+ parts
+
+ fun mk_P (MutualPart {cargTs, ...}) P =
+ let
+ val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
+ val atup = foldr1 HOLogic.mk_prod avars
+ in
+ tupled_lambda atup (list_comb (P, avars))
+ end
+
+ val Ps = map2 mk_P parts newPs
+ val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
+
+ val induct_inst =
+ forall_elim (cert case_exp) induct
+ |> full_simplify SumTree.sumcase_split_ss
+ |> full_simplify (HOL_basic_ss addsimps all_f_defs)
+
+ fun project rule (MutualPart {cargTs, i, ...}) k =
+ let
+ val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
+ val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
+ in
+ (rule
+ |> forall_elim (cert inj)
+ |> full_simplify SumTree.sumcase_split_ss
+ |> fold_rev (forall_intr o cert) (afs @ newPs),
+ k + length cargTs)
+ end
+ in
+ fst (fold_map (project induct_inst) parts 0)
+ end
+
+
+fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
+ let
+ val result = inner_cont proof
+ val FundefResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
+ termination,domintros} = result
+
+ val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
+ (mk_applied_form lthy cargTs (symmetric f_def), f))
+ parts
+ |> split_list
+
+ val all_orig_fdefs = map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
+
+ fun mk_mpsimp fqgar sum_psimp =
+ in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
+
+ val rew_ss = HOL_basic_ss addsimps all_f_defs
+ val mpsimps = map2 mk_mpsimp fqgars psimps
+ val mtrsimps = map_option (map2 mk_mpsimp fqgars) trsimps
+ val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
+ val mtermination = full_simplify rew_ss termination
+ val mdomintros = map_option (map (full_simplify rew_ss)) domintros
+ in
+ FundefResult { fs=fs, G=G, R=R,
+ psimps=mpsimps, simple_pinducts=minducts,
+ cases=cases, termination=mtermination,
+ domintros=mdomintros,
+ trsimps=mtrsimps}
+ end
+
+fun prepare_fundef_mutual config defname fixes eqss lthy =
+ let
+ val mutual = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
+ val Mutual {fsum_var=(n, T), qglrs, ...} = mutual
+
+ val ((fsum, goalstate, cont), lthy') =
+ FundefCore.prepare_fundef config defname [((n, T), NoSyn)] qglrs lthy
+
+ val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
+
+ val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual'
+ in
+ ((goalstate, mutual_cont), lthy'')
+ end
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/pattern_split.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,137 @@
+(* Title: HOL/Tools/Function/pattern_split.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+
+Automatic splitting of overlapping constructor patterns. This is a preprocessing step which
+turns a specification with overlaps into an overlap-free specification.
+
+*)
+
+signature FUNDEF_SPLIT =
+sig
+ val split_some_equations :
+ Proof.context -> (bool * term) list -> term list list
+
+ val split_all_equations :
+ Proof.context -> term list -> term list list
+end
+
+structure FundefSplit : FUNDEF_SPLIT =
+struct
+
+open FundefLib
+
+(* We use proof context for the variable management *)
+(* FIXME: no __ *)
+
+fun new_var ctx vs T =
+ let
+ val [v] = Variable.variant_frees ctx vs [("v", T)]
+ in
+ (Free v :: vs, Free v)
+ end
+
+fun saturate ctx vs t =
+ fold (fn T => fn (vs, t) => new_var ctx vs T |> apsnd (curry op $ t))
+ (binder_types (fastype_of t)) (vs, t)
+
+
+(* This is copied from "fundef_datatype.ML" *)
+fun inst_constrs_of thy (T as Type (name, _)) =
+ map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
+ (the (Datatype.get_datatype_constrs thy name))
+ | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
+
+
+
+
+fun join ((vs1,sub1), (vs2,sub2)) = (merge (op aconv) (vs1,vs2), sub1 @ sub2)
+fun join_product (xs, ys) = map_product (curry join) xs ys
+
+fun join_list [] = []
+ | join_list xs = foldr1 (join_product) xs
+
+
+exception DISJ
+
+fun pattern_subtract_subst ctx vs t t' =
+ let
+ exception DISJ
+ fun pattern_subtract_subst_aux vs _ (Free v2) = []
+ | pattern_subtract_subst_aux vs (v as (Free (_, T))) t' =
+ let
+ fun foo constr =
+ let
+ val (vs', t) = saturate ctx vs constr
+ val substs = pattern_subtract_subst ctx vs' t t'
+ in
+ map (fn (vs, subst) => (vs, (v,t)::subst)) substs
+ end
+ in
+ flat (map foo (inst_constrs_of (ProofContext.theory_of ctx) T))
+ end
+ | pattern_subtract_subst_aux vs t t' =
+ let
+ val (C, ps) = strip_comb t
+ val (C', qs) = strip_comb t'
+ in
+ if C = C'
+ then flat (map2 (pattern_subtract_subst_aux vs) ps qs)
+ else raise DISJ
+ end
+ in
+ pattern_subtract_subst_aux vs t t'
+ handle DISJ => [(vs, [])]
+ end
+
+
+(* p - q *)
+fun pattern_subtract ctx eq2 eq1 =
+ let
+ val thy = ProofContext.theory_of ctx
+
+ val (vs, feq1 as (_ $ (_ $ lhs1 $ _))) = dest_all_all eq1
+ val (_, _ $ (_ $ lhs2 $ _)) = dest_all_all eq2
+
+ val substs = pattern_subtract_subst ctx vs lhs1 lhs2
+
+ fun instantiate (vs', sigma) =
+ let
+ val t = Pattern.rewrite_term thy sigma [] feq1
+ in
+ fold_rev Logic.all (map Free (frees_in_term ctx t) inter vs') t
+ end
+ in
+ map instantiate substs
+ end
+
+
+(* ps - p' *)
+fun pattern_subtract_from_many ctx p'=
+ flat o map (pattern_subtract ctx p')
+
+(* in reverse order *)
+fun pattern_subtract_many ctx ps' =
+ fold_rev (pattern_subtract_from_many ctx) ps'
+
+
+
+fun split_some_equations ctx eqns =
+ let
+ fun split_aux prev [] = []
+ | split_aux prev ((true, eq) :: es) = pattern_subtract_many ctx prev [eq]
+ :: split_aux (eq :: prev) es
+ | split_aux prev ((false, eq) :: es) = [eq]
+ :: split_aux (eq :: prev) es
+ in
+ split_aux [] eqns
+ end
+
+fun split_all_equations ctx =
+ split_some_equations ctx o map (pair true)
+
+
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,429 @@
+(* Title: HOL/Tools/Function/scnp_reconstruct.ML
+ Author: Armin Heller, TU Muenchen
+ Author: Alexander Krauss, TU Muenchen
+
+Proof reconstruction for SCNP
+*)
+
+signature SCNP_RECONSTRUCT =
+sig
+
+ val sizechange_tac : Proof.context -> tactic -> tactic
+
+ val decomp_scnp : ScnpSolve.label list -> Proof.context -> Proof.method
+
+ val setup : theory -> theory
+
+ datatype multiset_setup =
+ Multiset of
+ {
+ msetT : typ -> typ,
+ mk_mset : typ -> term list -> term,
+ mset_regroup_conv : int list -> conv,
+ mset_member_tac : int -> int -> tactic,
+ mset_nonempty_tac : int -> tactic,
+ mset_pwleq_tac : int -> tactic,
+ set_of_simps : thm list,
+ smsI' : thm,
+ wmsI2'' : thm,
+ wmsI1 : thm,
+ reduction_pair : thm
+ }
+
+
+ val multiset_setup : multiset_setup -> theory -> theory
+
+end
+
+structure ScnpReconstruct : SCNP_RECONSTRUCT =
+struct
+
+val PROFILE = FundefCommon.PROFILE
+fun TRACE x = if ! FundefCommon.profile then Output.tracing x else ()
+
+open ScnpSolve
+
+val natT = HOLogic.natT
+val nat_pairT = HOLogic.mk_prodT (natT, natT)
+
+(* Theory dependencies *)
+
+datatype multiset_setup =
+ Multiset of
+ {
+ msetT : typ -> typ,
+ mk_mset : typ -> term list -> term,
+ mset_regroup_conv : int list -> conv,
+ mset_member_tac : int -> int -> tactic,
+ mset_nonempty_tac : int -> tactic,
+ mset_pwleq_tac : int -> tactic,
+ set_of_simps : thm list,
+ smsI' : thm,
+ wmsI2'' : thm,
+ wmsI1 : thm,
+ reduction_pair : thm
+ }
+
+structure MultisetSetup = TheoryDataFun
+(
+ type T = multiset_setup option
+ val empty = NONE
+ val copy = I;
+ val extend = I;
+ fun merge _ (v1, v2) = if is_some v2 then v2 else v1
+)
+
+val multiset_setup = MultisetSetup.put o SOME
+
+fun undef x = error "undef"
+fun get_multiset_setup thy = MultisetSetup.get thy
+ |> the_default (Multiset
+{ msetT = undef, mk_mset=undef,
+ mset_regroup_conv=undef, mset_member_tac = undef,
+ mset_nonempty_tac = undef, mset_pwleq_tac = undef,
+ set_of_simps = [],reduction_pair = refl,
+ smsI'=refl, wmsI2''=refl, wmsI1=refl })
+
+fun order_rpair _ MAX = @{thm max_rpair_set}
+ | order_rpair msrp MS = msrp
+ | order_rpair _ MIN = @{thm min_rpair_set}
+
+fun ord_intros_max true =
+ (@{thm smax_emptyI}, @{thm smax_insertI})
+ | ord_intros_max false =
+ (@{thm wmax_emptyI}, @{thm wmax_insertI})
+fun ord_intros_min true =
+ (@{thm smin_emptyI}, @{thm smin_insertI})
+ | ord_intros_min false =
+ (@{thm wmin_emptyI}, @{thm wmin_insertI})
+
+fun gen_probl D cs =
+ let
+ val n = Termination.get_num_points D
+ val arity = length o Termination.get_measures D
+ fun measure p i = nth (Termination.get_measures D p) i
+
+ fun mk_graph c =
+ let
+ val (_, p, _, q, _, _) = Termination.dest_call D c
+
+ fun add_edge i j =
+ case Termination.get_descent D c (measure p i) (measure q j)
+ of SOME (Termination.Less _) => cons (i, GTR, j)
+ | SOME (Termination.LessEq _) => cons (i, GEQ, j)
+ | _ => I
+
+ val edges =
+ fold_product add_edge (0 upto arity p - 1) (0 upto arity q - 1) []
+ in
+ G (p, q, edges)
+ end
+ in
+ GP (map arity (0 upto n - 1), map mk_graph cs)
+ end
+
+(* General reduction pair application *)
+fun rem_inv_img ctxt =
+ let
+ val unfold_tac = LocalDefs.unfold_tac ctxt
+ in
+ rtac @{thm subsetI} 1
+ THEN etac @{thm CollectE} 1
+ THEN REPEAT (etac @{thm exE} 1)
+ THEN unfold_tac @{thms inv_image_def}
+ THEN rtac @{thm CollectI} 1
+ THEN etac @{thm conjE} 1
+ THEN etac @{thm ssubst} 1
+ THEN unfold_tac (@{thms split_conv} @ @{thms triv_forall_equality}
+ @ @{thms sum.cases})
+ end
+
+(* Sets *)
+
+val setT = HOLogic.mk_setT
+
+fun set_member_tac m i =
+ if m = 0 then rtac @{thm insertI1} i
+ else rtac @{thm insertI2} i THEN set_member_tac (m - 1) i
+
+val set_nonempty_tac = rtac @{thm insert_not_empty}
+
+fun set_finite_tac i =
+ rtac @{thm finite.emptyI} i
+ ORELSE (rtac @{thm finite.insertI} i THEN (fn st => set_finite_tac i st))
+
+
+(* Reconstruction *)
+
+fun reconstruct_tac ctxt D cs (gp as GP (_, gs)) certificate =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val Multiset
+ { msetT, mk_mset,
+ mset_regroup_conv, mset_member_tac,
+ mset_nonempty_tac, mset_pwleq_tac, set_of_simps,
+ smsI', wmsI2'', wmsI1, reduction_pair=ms_rp }
+ = get_multiset_setup thy
+
+ fun measure_fn p = nth (Termination.get_measures D p)
+
+ fun get_desc_thm cidx m1 m2 bStrict =
+ case Termination.get_descent D (nth cs cidx) m1 m2
+ of SOME (Termination.Less thm) =>
+ if bStrict then thm
+ else (thm COMP (Thm.lift_rule (cprop_of thm) @{thm less_imp_le}))
+ | SOME (Termination.LessEq (thm, _)) =>
+ if not bStrict then thm
+ else sys_error "get_desc_thm"
+ | _ => sys_error "get_desc_thm"
+
+ val (label, lev, sl, covering) = certificate
+
+ fun prove_lev strict g =
+ let
+ val G (p, q, el) = nth gs g
+
+ fun less_proof strict (j, b) (i, a) =
+ let
+ val tag_flag = b < a orelse (not strict andalso b <= a)
+
+ val stored_thm =
+ get_desc_thm g (measure_fn p i) (measure_fn q j)
+ (not tag_flag)
+ |> Conv.fconv_rule (Thm.beta_conversion true)
+
+ val rule = if strict
+ then if b < a then @{thm pair_lessI2} else @{thm pair_lessI1}
+ else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
+ in
+ rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
+ THEN (if tag_flag then Arith_Data.verbose_arith_tac ctxt 1 else all_tac)
+ end
+
+ fun steps_tac MAX strict lq lp =
+ let
+ val (empty, step) = ord_intros_max strict
+ in
+ if length lq = 0
+ then rtac empty 1 THEN set_finite_tac 1
+ THEN (if strict then set_nonempty_tac 1 else all_tac)
+ else
+ let
+ val (j, b) :: rest = lq
+ val (i, a) = the (covering g strict j)
+ fun choose xs = set_member_tac (Library.find_index (curry op = (i, a)) xs) 1
+ val solve_tac = choose lp THEN less_proof strict (j, b) (i, a)
+ in
+ rtac step 1 THEN solve_tac THEN steps_tac MAX strict rest lp
+ end
+ end
+ | steps_tac MIN strict lq lp =
+ let
+ val (empty, step) = ord_intros_min strict
+ in
+ if length lp = 0
+ then rtac empty 1
+ THEN (if strict then set_nonempty_tac 1 else all_tac)
+ else
+ let
+ val (i, a) :: rest = lp
+ val (j, b) = the (covering g strict i)
+ fun choose xs = set_member_tac (Library.find_index (curry op = (j, b)) xs) 1
+ val solve_tac = choose lq THEN less_proof strict (j, b) (i, a)
+ in
+ rtac step 1 THEN solve_tac THEN steps_tac MIN strict lq rest
+ end
+ end
+ | steps_tac MS strict lq lp =
+ let
+ fun get_str_cover (j, b) =
+ if is_some (covering g true j) then SOME (j, b) else NONE
+ fun get_wk_cover (j, b) = the (covering g false j)
+
+ val qs = lq \\ map_filter get_str_cover lq
+ val ps = map get_wk_cover qs
+
+ fun indices xs ys = map (fn y => Library.find_index (curry op = y) xs) ys
+ val iqs = indices lq qs
+ val ips = indices lp ps
+
+ local open Conv in
+ fun t_conv a C =
+ params_conv ~1 (K ((concl_conv ~1 o arg_conv o arg1_conv o a) C)) ctxt
+ val goal_rewrite =
+ t_conv arg1_conv (mset_regroup_conv iqs)
+ then_conv t_conv arg_conv (mset_regroup_conv ips)
+ end
+ in
+ CONVERSION goal_rewrite 1
+ THEN (if strict then rtac smsI' 1
+ else if qs = lq then rtac wmsI2'' 1
+ else rtac wmsI1 1)
+ THEN mset_pwleq_tac 1
+ THEN EVERY (map2 (less_proof false) qs ps)
+ THEN (if strict orelse qs <> lq
+ then LocalDefs.unfold_tac ctxt set_of_simps
+ THEN steps_tac MAX true (lq \\ qs) (lp \\ ps)
+ else all_tac)
+ end
+ in
+ rem_inv_img ctxt
+ THEN steps_tac label strict (nth lev q) (nth lev p)
+ end
+
+ val (mk_set, setT) = if label = MS then (mk_mset, msetT) else (HOLogic.mk_set, setT)
+
+ fun tag_pair p (i, tag) =
+ HOLogic.pair_const natT natT $
+ (measure_fn p i $ Bound 0) $ HOLogic.mk_number natT tag
+
+ fun pt_lev (p, lm) = Abs ("x", Termination.get_types D p,
+ mk_set nat_pairT (map (tag_pair p) lm))
+
+ val level_mapping =
+ map_index pt_lev lev
+ |> Termination.mk_sumcases D (setT nat_pairT)
+ |> cterm_of thy
+ in
+ PROFILE "Proof Reconstruction"
+ (CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv sl))) 1
+ THEN (rtac @{thm reduction_pair_lemma} 1)
+ THEN (rtac @{thm rp_inv_image_rp} 1)
+ THEN (rtac (order_rpair ms_rp label) 1)
+ THEN PRIMITIVE (instantiate' [] [SOME level_mapping])
+ THEN unfold_tac @{thms rp_inv_image_def} (local_simpset_of ctxt)
+ THEN LocalDefs.unfold_tac ctxt
+ (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv})
+ THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}]))
+ THEN EVERY (map (prove_lev true) sl)
+ THEN EVERY (map (prove_lev false) ((0 upto length cs - 1) \\ sl)))
+ end
+
+
+
+local open Termination in
+fun print_cell (SOME (Less _)) = "<"
+ | print_cell (SOME (LessEq _)) = "\<le>"
+ | print_cell (SOME (None _)) = "-"
+ | print_cell (SOME (False _)) = "-"
+ | print_cell (NONE) = "?"
+
+fun print_error ctxt D = CALLS (fn (cs, i) =>
+ let
+ val np = get_num_points D
+ val ms = map (get_measures D) (0 upto np - 1)
+ val tys = map (get_types D) (0 upto np - 1)
+ fun index xs = (1 upto length xs) ~~ xs
+ fun outp s t f xs = map (fn (x, y) => s ^ Int.toString x ^ t ^ f y ^ "\n") xs
+ val ims = index (map index ms)
+ val _ = Output.tracing (concat (outp "fn #" ":\n" (concat o outp "\tmeasure #" ": " (Syntax.string_of_term ctxt)) ims))
+ fun print_call (k, c) =
+ let
+ val (_, p, _, q, _, _) = dest_call D c
+ val _ = Output.tracing ("call table for call #" ^ Int.toString k ^ ": fn " ^
+ Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1))
+ val caller_ms = nth ms p
+ val callee_ms = nth ms q
+ val entries = map (fn x => map (pair x) (callee_ms)) (caller_ms)
+ fun print_ln (i : int, l) = concat (Int.toString i :: " " :: map (enclose " " " " o print_cell o (uncurry (get_descent D c))) l)
+ val _ = Output.tracing (concat (Int.toString (p + 1) ^ "|" ^ Int.toString (q + 1) ^
+ " " :: map (enclose " " " " o Int.toString) (1 upto length callee_ms)) ^ "\n"
+ ^ cat_lines (map print_ln ((1 upto (length entries)) ~~ entries)))
+ in
+ true
+ end
+ fun list_call (k, c) =
+ let
+ val (_, p, _, q, _, _) = dest_call D c
+ val _ = Output.tracing ("call #" ^ (Int.toString k) ^ ": fn " ^
+ Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1) ^ "\n" ^
+ (Syntax.string_of_term ctxt c))
+ in true end
+ val _ = forall list_call ((1 upto length cs) ~~ cs)
+ val _ = forall print_call ((1 upto length cs) ~~ cs)
+ in
+ all_tac
+ end)
+end
+
+
+fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
+ let
+ val ms_configured = is_some (MultisetSetup.get (ProofContext.theory_of ctxt))
+ val orders' = if ms_configured then orders
+ else filter_out (curry op = MS) orders
+ val gp = gen_probl D cs
+(* val _ = TRACE ("SCNP instance: " ^ makestring gp)*)
+ val certificate = generate_certificate use_tags orders' gp
+(* val _ = TRACE ("Certificate: " ^ makestring certificate)*)
+
+ in
+ case certificate
+ of NONE => err_cont D i
+ | SOME cert =>
+ SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i
+ THEN (rtac @{thm wf_empty} i ORELSE cont D i)
+ end)
+
+fun gen_decomp_scnp_tac orders autom_tac ctxt err_cont =
+ let
+ open Termination
+ val derive_diag = Descent.derive_diag ctxt autom_tac
+ val derive_all = Descent.derive_all ctxt autom_tac
+ val decompose = Decompose.decompose_tac ctxt autom_tac
+ val scnp_no_tags = single_scnp_tac false orders ctxt
+ val scnp_full = single_scnp_tac true orders ctxt
+
+ fun first_round c e =
+ derive_diag (REPEAT scnp_no_tags c e)
+
+ val second_round =
+ REPEAT (fn c => fn e => decompose (scnp_no_tags c c) e)
+
+ val third_round =
+ derive_all oo
+ REPEAT (fn c => fn e =>
+ scnp_full (decompose c c) e)
+
+ fun Then s1 s2 c e = s1 (s2 c c) (s2 c e)
+
+ val strategy = Then (Then first_round second_round) third_round
+
+ in
+ TERMINATION ctxt (strategy err_cont err_cont)
+ end
+
+fun gen_sizechange_tac orders autom_tac ctxt err_cont =
+ TRY (FundefCommon.apply_termination_rule ctxt 1)
+ THEN TRY (Termination.wf_union_tac ctxt)
+ THEN
+ (rtac @{thm wf_empty} 1
+ ORELSE gen_decomp_scnp_tac orders autom_tac ctxt err_cont 1)
+
+fun sizechange_tac ctxt autom_tac =
+ gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt (K (K all_tac))
+
+fun decomp_scnp orders ctxt =
+ let
+ val extra_simps = FundefCommon.TerminationSimps.get ctxt
+ val autom_tac = auto_tac (local_clasimpset_of ctxt addsimps2 extra_simps)
+ in
+ SIMPLE_METHOD
+ (gen_sizechange_tac orders autom_tac ctxt (print_error ctxt))
+ end
+
+
+(* Method setup *)
+
+val orders =
+ Scan.repeat1
+ ((Args.$$$ "max" >> K MAX) ||
+ (Args.$$$ "min" >> K MIN) ||
+ (Args.$$$ "ms" >> K MS))
+ || Scan.succeed [MAX, MS, MIN]
+
+val setup = Method.setup @{binding sizechange}
+ (Scan.lift orders --| Method.sections clasimp_modifiers >> decomp_scnp)
+ "termination prover with graph decomposition and the NP subset of size change termination"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/scnp_solve.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,257 @@
+(* Title: HOL/Tools/Function/scnp_solve.ML
+ Author: Armin Heller, TU Muenchen
+ Author: Alexander Krauss, TU Muenchen
+
+Generate certificates for SCNP using a SAT solver
+*)
+
+
+signature SCNP_SOLVE =
+sig
+
+ datatype edge = GTR | GEQ
+ datatype graph = G of int * int * (int * edge * int) list
+ datatype graph_problem = GP of int list * graph list
+
+ datatype label = MIN | MAX | MS
+
+ type certificate =
+ label (* which order *)
+ * (int * int) list list (* (multi)sets *)
+ * int list (* strictly ordered calls *)
+ * (int -> bool -> int -> (int * int) option) (* covering function *)
+
+ val generate_certificate : bool -> label list -> graph_problem -> certificate option
+
+ val solver : string ref
+end
+
+structure ScnpSolve : SCNP_SOLVE =
+struct
+
+(** Graph problems **)
+
+datatype edge = GTR | GEQ ;
+datatype graph = G of int * int * (int * edge * int) list ;
+datatype graph_problem = GP of int list * graph list ;
+
+datatype label = MIN | MAX | MS ;
+type certificate =
+ label
+ * (int * int) list list
+ * int list
+ * (int -> bool -> int -> (int * int) option)
+
+fun graph_at (GP (_, gs), i) = nth gs i ;
+fun num_prog_pts (GP (arities, _)) = length arities ;
+fun num_graphs (GP (_, gs)) = length gs ;
+fun arity (GP (arities, gl)) i = nth arities i ;
+fun ndigits (GP (arities, _)) = IntInf.log2 (List.foldl (op +) 0 arities) + 1
+
+
+(** Propositional formulas **)
+
+val Not = PropLogic.Not and And = PropLogic.And and Or = PropLogic.Or
+val BoolVar = PropLogic.BoolVar
+fun Implies (p, q) = Or (Not p, q)
+fun Equiv (p, q) = And (Implies (p, q), Implies (q, p))
+val all = PropLogic.all
+
+(* finite indexed quantifiers:
+
+iforall n f <==> /\
+ / \ f i
+ 0<=i<n
+ *)
+fun iforall n f = all (map f (0 upto n - 1))
+fun iexists n f = PropLogic.exists (map f (0 upto n - 1))
+fun iforall2 n m f = all (map_product f (0 upto n - 1) (0 upto m - 1))
+
+fun the_one var n x = all (var x :: map (Not o var) ((0 upto n - 1) \\ [x]))
+fun exactly_one n f = iexists n (the_one f n)
+
+(* SAT solving *)
+val solver = ref "auto";
+fun sat_solver x =
+ FundefCommon.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
+
+(* "Virtual constructors" for various propositional variables *)
+fun var_constrs (gp as GP (arities, gl)) =
+ let
+ val n = Int.max (num_graphs gp, num_prog_pts gp)
+ val k = List.foldl Int.max 1 arities
+
+ (* Injective, provided a < 8, x < n, and i < k. *)
+ fun prod a x i j = ((j * k + i) * n + x) * 8 + a + 1
+
+ fun ES (g, i, j) = BoolVar (prod 0 g i j)
+ fun EW (g, i, j) = BoolVar (prod 1 g i j)
+ fun WEAK g = BoolVar (prod 2 g 0 0)
+ fun STRICT g = BoolVar (prod 3 g 0 0)
+ fun P (p, i) = BoolVar (prod 4 p i 0)
+ fun GAM (g, i, j)= BoolVar (prod 5 g i j)
+ fun EPS (g, i) = BoolVar (prod 6 g i 0)
+ fun TAG (p, i) b = BoolVar (prod 7 p i b)
+ in
+ (ES,EW,WEAK,STRICT,P,GAM,EPS,TAG)
+ end
+
+
+fun graph_info gp g =
+ let
+ val G (p, q, edgs) = graph_at (gp, g)
+ in
+ (g, p, q, arity gp p, arity gp q, edgs)
+ end
+
+
+(* Order-independent part of encoding *)
+
+fun encode_graphs bits gp =
+ let
+ val ng = num_graphs gp
+ val (ES,EW,_,_,_,_,_,TAG) = var_constrs gp
+
+ fun encode_constraint_strict 0 (x, y) = PropLogic.False
+ | encode_constraint_strict k (x, y) =
+ Or (And (TAG x (k - 1), Not (TAG y (k - 1))),
+ And (Equiv (TAG x (k - 1), TAG y (k - 1)),
+ encode_constraint_strict (k - 1) (x, y)))
+
+ fun encode_constraint_weak k (x, y) =
+ Or (encode_constraint_strict k (x, y),
+ iforall k (fn i => Equiv (TAG x i, TAG y i)))
+
+ fun encode_graph (g, p, q, n, m, edges) =
+ let
+ fun encode_edge i j =
+ if exists (fn x => x = (i, GTR, j)) edges then
+ And (ES (g, i, j), EW (g, i, j))
+ else if not (exists (fn x => x = (i, GEQ, j)) edges) then
+ And (Not (ES (g, i, j)), Not (EW (g, i, j)))
+ else
+ And (
+ Equiv (ES (g, i, j),
+ encode_constraint_strict bits ((p, i), (q, j))),
+ Equiv (EW (g, i, j),
+ encode_constraint_weak bits ((p, i), (q, j))))
+ in
+ iforall2 n m encode_edge
+ end
+ in
+ iforall ng (encode_graph o graph_info gp)
+ end
+
+
+(* Order-specific part of encoding *)
+
+fun encode bits gp mu =
+ let
+ val ng = num_graphs gp
+ val (ES,EW,WEAK,STRICT,P,GAM,EPS,_) = var_constrs gp
+
+ fun encode_graph MAX (g, p, q, n, m, _) =
+ And (
+ Equiv (WEAK g,
+ iforall m (fn j =>
+ Implies (P (q, j),
+ iexists n (fn i =>
+ And (P (p, i), EW (g, i, j)))))),
+ Equiv (STRICT g,
+ And (
+ iforall m (fn j =>
+ Implies (P (q, j),
+ iexists n (fn i =>
+ And (P (p, i), ES (g, i, j))))),
+ iexists n (fn i => P (p, i)))))
+ | encode_graph MIN (g, p, q, n, m, _) =
+ And (
+ Equiv (WEAK g,
+ iforall n (fn i =>
+ Implies (P (p, i),
+ iexists m (fn j =>
+ And (P (q, j), EW (g, i, j)))))),
+ Equiv (STRICT g,
+ And (
+ iforall n (fn i =>
+ Implies (P (p, i),
+ iexists m (fn j =>
+ And (P (q, j), ES (g, i, j))))),
+ iexists m (fn j => P (q, j)))))
+ | encode_graph MS (g, p, q, n, m, _) =
+ all [
+ Equiv (WEAK g,
+ iforall m (fn j =>
+ Implies (P (q, j),
+ iexists n (fn i => GAM (g, i, j))))),
+ Equiv (STRICT g,
+ iexists n (fn i =>
+ And (P (p, i), Not (EPS (g, i))))),
+ iforall2 n m (fn i => fn j =>
+ Implies (GAM (g, i, j),
+ all [
+ P (p, i),
+ P (q, j),
+ EW (g, i, j),
+ Equiv (Not (EPS (g, i)), ES (g, i, j))])),
+ iforall n (fn i =>
+ Implies (And (P (p, i), EPS (g, i)),
+ exactly_one m (fn j => GAM (g, i, j))))
+ ]
+ in
+ all [
+ encode_graphs bits gp,
+ iforall ng (encode_graph mu o graph_info gp),
+ iforall ng (fn x => WEAK x),
+ iexists ng (fn x => STRICT x)
+ ]
+ end
+
+
+(*Generieren des level-mapping und diverser output*)
+fun mk_certificate bits label gp f =
+ let
+ val (ES,EW,WEAK,STRICT,P,GAM,EPS,TAG) = var_constrs gp
+ fun assign (PropLogic.BoolVar v) = the_default false (f v)
+ fun assignTag i j =
+ (fold (fn x => fn y => 2 * y + (if assign (TAG (i, j) x) then 1 else 0))
+ (bits - 1 downto 0) 0)
+
+ val level_mapping =
+ let fun prog_pt_mapping p =
+ map_filter (fn x => if assign (P(p, x)) then SOME (x, assignTag p x) else NONE)
+ (0 upto (arity gp p) - 1)
+ in map prog_pt_mapping (0 upto num_prog_pts gp - 1) end
+
+ val strict_list = filter (assign o STRICT) (0 upto num_graphs gp - 1)
+
+ fun covering_pair g bStrict j =
+ let
+ val (_, p, q, n, m, _) = graph_info gp g
+
+ fun cover MAX j = find_index (fn i => assign (P (p, i)) andalso assign (EW (g, i, j))) (0 upto n - 1)
+ | cover MS k = find_index (fn i => assign (GAM (g, i, k))) (0 upto n - 1)
+ | cover MIN i = find_index (fn j => assign (P (q, j)) andalso assign (EW (g, i, j))) (0 upto m - 1)
+ fun cover_strict MAX j = find_index (fn i => assign (P (p, i)) andalso assign (ES (g, i, j))) (0 upto n - 1)
+ | cover_strict MS k = find_index (fn i => assign (GAM (g, i, k)) andalso not (assign (EPS (g, i) ))) (0 upto n - 1)
+ | cover_strict MIN i = find_index (fn j => assign (P (q, j)) andalso assign (ES (g, i, j))) (0 upto m - 1)
+ val i = if bStrict then cover_strict label j else cover label j
+ in
+ find_first (fn x => fst x = i) (nth level_mapping (if label = MIN then q else p))
+ end
+ in
+ (label, level_mapping, strict_list, covering_pair)
+ end
+
+(*interface for the proof reconstruction*)
+fun generate_certificate use_tags labels gp =
+ let
+ val bits = if use_tags then ndigits gp else 0
+ in
+ get_first
+ (fn l => case sat_solver (encode bits gp l) of
+ SatSolver.SATISFIABLE f => SOME (mk_certificate bits l gp f)
+ | _ => NONE)
+ labels
+ end
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/size.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,242 @@
+(* Title: HOL/Tools/Function/size.ML
+ Author: Stefan Berghofer, Florian Haftmann & Alexander Krauss, TU Muenchen
+
+Size functions for datatypes.
+*)
+
+signature SIZE =
+sig
+ val size_thms: theory -> string -> thm list
+ val setup: theory -> theory
+end;
+
+structure Size: SIZE =
+struct
+
+open DatatypeAux;
+
+structure SizeData = TheoryDataFun
+(
+ type T = (string * thm list) Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I
+ val extend = I
+ fun merge _ = Symtab.merge (K true);
+);
+
+val lookup_size = SizeData.get #> Symtab.lookup;
+
+fun plus (t1, t2) = Const ("HOL.plus_class.plus",
+ HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2;
+
+fun size_of_type f g h (T as Type (s, Ts)) =
+ (case f s of
+ SOME t => SOME t
+ | NONE => (case g s of
+ SOME size_name =>
+ SOME (list_comb (Const (size_name,
+ map (fn U => U --> HOLogic.natT) Ts @ [T] ---> HOLogic.natT),
+ map (size_of_type' f g h) Ts))
+ | NONE => NONE))
+ | size_of_type f g h (TFree (s, _)) = h s
+and size_of_type' f g h T = (case size_of_type f g h T of
+ NONE => Abs ("x", T, HOLogic.zero)
+ | SOME t => t);
+
+fun is_poly thy (DtType (name, dts)) =
+ (case Datatype.get_datatype thy name of
+ NONE => false
+ | SOME _ => exists (is_poly thy) dts)
+ | is_poly _ _ = true;
+
+fun constrs_of thy name =
+ let
+ val {descr, index, ...} = Datatype.the_datatype thy name
+ val SOME (_, _, constrs) = AList.lookup op = descr index
+ in constrs end;
+
+val app = curry (list_comb o swap);
+
+fun prove_size_thms (info : info) new_type_names thy =
+ let
+ val {descr, alt_names, sorts, rec_names, rec_rewrites, induction, ...} = info;
+ val l = length new_type_names;
+ val alt_names' = (case alt_names of
+ NONE => replicate l NONE | SOME names => map SOME names);
+ val descr' = List.take (descr, l);
+ val (rec_names1, rec_names2) = chop l rec_names;
+ val recTs = get_rec_types descr sorts;
+ val (recTs1, recTs2) = chop l recTs;
+ val (_, (_, paramdts, _)) :: _ = descr;
+ val paramTs = map (typ_of_dtyp descr sorts) paramdts;
+ val ((param_size_fs, param_size_fTs), f_names) = paramTs |>
+ map (fn T as TFree (s, _) =>
+ let
+ val name = "f" ^ implode (tl (explode s));
+ val U = T --> HOLogic.natT
+ in
+ (((s, Free (name, U)), U), name)
+ end) |> split_list |>> split_list;
+ val param_size = AList.lookup op = param_size_fs;
+
+ val extra_rewrites = descr |> map (#1 o snd) |> distinct op = |>
+ map_filter (Option.map snd o lookup_size thy) |> flat;
+ val extra_size = Option.map fst o lookup_size thy;
+
+ val (((size_names, size_fns), def_names), def_names') =
+ recTs1 ~~ alt_names' |>
+ map (fn (T as Type (s, _), optname) =>
+ let
+ val s' = the_default (Long_Name.base_name s) optname ^ "_size";
+ val s'' = Sign.full_bname thy s'
+ in
+ (s'',
+ (list_comb (Const (s'', param_size_fTs @ [T] ---> HOLogic.natT),
+ map snd param_size_fs),
+ (s' ^ "_def", s' ^ "_overloaded_def")))
+ end) |> split_list ||>> split_list ||>> split_list;
+ val overloaded_size_fns = map HOLogic.size_const recTs1;
+
+ (* instantiation for primrec combinator *)
+ fun size_of_constr b size_ofp ((_, cargs), (_, cargs')) =
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val k = length (filter is_rec_type cargs);
+ val (ts, _, _) = fold_rev (fn ((dt, dt'), T) => fn (us, i, j) =>
+ if is_rec_type dt then (Bound i :: us, i + 1, j + 1)
+ else
+ (if b andalso is_poly thy dt' then
+ case size_of_type (K NONE) extra_size size_ofp T of
+ NONE => us | SOME sz => sz $ Bound j :: us
+ else us, i, j + 1))
+ (cargs ~~ cargs' ~~ Ts) ([], 0, k);
+ val t =
+ if null ts andalso (not b orelse not (exists (is_poly thy) cargs'))
+ then HOLogic.zero
+ else foldl1 plus (ts @ [HOLogic.Suc_zero])
+ in
+ List.foldr (fn (T, t') => Abs ("x", T, t')) t (Ts @ replicate k HOLogic.natT)
+ end;
+
+ val fs = maps (fn (_, (name, _, constrs)) =>
+ map (size_of_constr true param_size) (constrs ~~ constrs_of thy name)) descr;
+ val fs' = maps (fn (n, (name, _, constrs)) =>
+ map (size_of_constr (l <= n) (K NONE)) (constrs ~~ constrs_of thy name)) descr;
+ val fTs = map fastype_of fs;
+
+ val (rec_combs1, rec_combs2) = chop l (map (fn (T, rec_name) =>
+ Const (rec_name, fTs @ [T] ---> HOLogic.natT))
+ (recTs ~~ rec_names));
+
+ fun define_overloaded (def_name, eq) lthy =
+ let
+ val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq;
+ val ((_, (_, thm)), lthy') = lthy |> LocalTheory.define Thm.definitionK
+ ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs));
+ val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy');
+ val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
+ in (thm', lthy') end;
+
+ val ((size_def_thms, size_def_thms'), thy') =
+ thy
+ |> Sign.add_consts_i (map (fn (s, T) =>
+ (Binding.name (Long_Name.base_name s), param_size_fTs @ [T] ---> HOLogic.natT, NoSyn))
+ (size_names ~~ recTs1))
+ |> PureThy.add_defs false
+ (map (Thm.no_attributes o apsnd (Logic.mk_equals o apsnd (app fs)))
+ (map Binding.name def_names ~~ (size_fns ~~ rec_combs1)))
+ ||> TheoryTarget.instantiation
+ (map (#1 o snd) descr', map dest_TFree paramTs, [HOLogic.class_size])
+ ||>> fold_map define_overloaded
+ (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1))
+ ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
+ ||> LocalTheory.exit_global;
+
+ val ctxt = ProofContext.init thy';
+
+ val simpset1 = HOL_basic_ss addsimps @{thm add_0} :: @{thm add_0_right} ::
+ size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites;
+ val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2);
+
+ fun mk_unfolded_size_eq tab size_ofp fs (p as (x, T), r) =
+ HOLogic.mk_eq (app fs r $ Free p,
+ the (size_of_type tab extra_size size_ofp T) $ Free p);
+
+ fun prove_unfolded_size_eqs size_ofp fs =
+ if null recTs2 then []
+ else split_conj_thm (SkipProof.prove ctxt xs []
+ (HOLogic.mk_Trueprop (mk_conj (replicate l HOLogic.true_const @
+ map (mk_unfolded_size_eq (AList.lookup op =
+ (new_type_names ~~ map (app fs) rec_combs1)) size_ofp fs)
+ (xs ~~ recTs2 ~~ rec_combs2))))
+ (fn _ => (indtac induction xs THEN_ALL_NEW asm_simp_tac simpset1) 1));
+
+ val unfolded_size_eqs1 = prove_unfolded_size_eqs param_size fs;
+ val unfolded_size_eqs2 = prove_unfolded_size_eqs (K NONE) fs';
+
+ (* characteristic equations for size functions *)
+ fun gen_mk_size_eq p size_of size_ofp size_const T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val tnames = Name.variant_list f_names (DatatypeProp.make_tnames Ts);
+ val ts = map_filter (fn (sT as (s, T), dt) =>
+ Option.map (fn sz => sz $ Free sT)
+ (if p dt then size_of_type size_of extra_size size_ofp T
+ else NONE)) (tnames ~~ Ts ~~ cargs)
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (size_const $ list_comb (Const (cname, Ts ---> T),
+ map2 (curry Free) tnames Ts),
+ if null ts then HOLogic.zero
+ else foldl1 plus (ts @ [HOLogic.Suc_zero])))
+ end;
+
+ val simpset2 = HOL_basic_ss addsimps
+ rec_rewrites @ size_def_thms @ unfolded_size_eqs1;
+ val simpset3 = HOL_basic_ss addsimps
+ rec_rewrites @ size_def_thms' @ unfolded_size_eqs2;
+
+ fun prove_size_eqs p size_fns size_ofp simpset =
+ maps (fn (((_, (_, _, constrs)), size_const), T) =>
+ map (fn constr => standard (SkipProof.prove ctxt [] []
+ (gen_mk_size_eq p (AList.lookup op = (new_type_names ~~ size_fns))
+ size_ofp size_const T constr)
+ (fn _ => simp_tac simpset 1))) constrs)
+ (descr' ~~ size_fns ~~ recTs1);
+
+ val size_eqns = prove_size_eqs (is_poly thy') size_fns param_size simpset2 @
+ prove_size_eqs is_rec_type overloaded_size_fns (K NONE) simpset3;
+
+ val ([size_thms], thy'') = PureThy.add_thmss
+ [((Binding.name "size", size_eqns),
+ [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
+ Thm.declaration_attribute
+ (fn thm => Context.mapping (Code.add_default_eqn thm) I)])] thy'
+
+ in
+ SizeData.map (fold (Symtab.update_new o apsnd (rpair size_thms))
+ (new_type_names ~~ size_names)) thy''
+ end;
+
+fun add_size_thms config (new_type_names as name :: _) thy =
+ let
+ val info as {descr, alt_names, ...} = Datatype.the_datatype thy name;
+ val prefix = Long_Name.map_base_name (K (space_implode "_"
+ (the_default (map Long_Name.base_name new_type_names) alt_names))) name;
+ val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt =>
+ is_rec_type dt andalso not (null (fst (strip_dtyp dt)))) cargs) constrs) descr
+ in if no_size then thy
+ else
+ thy
+ |> Sign.root_path
+ |> Sign.add_path prefix
+ |> Theory.checkpoint
+ |> prove_size_thms info new_type_names
+ |> Sign.restore_naming thy
+ end;
+
+val size_thms = snd oo (the oo lookup_size);
+
+val setup = Datatype.interpretation add_size_thms;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/sum_tree.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,43 @@
+(* Title: HOL/Tools/Function/sum_tree.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Some common tools for working with sum types in balanced tree form.
+*)
+
+structure SumTree =
+struct
+
+(* Theory dependencies *)
+val proj_in_rules = [@{thm "Datatype.Projl_Inl"}, @{thm "Datatype.Projr_Inr"}]
+val sumcase_split_ss = HOL_basic_ss addsimps (@{thm "Product_Type.split"} :: @{thms "sum.cases"})
+
+(* top-down access in balanced tree *)
+fun access_top_down {left, right, init} len i =
+ BalancedTree.access {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
+
+(* Sum types *)
+fun mk_sumT LT RT = Type ("+", [LT, RT])
+fun mk_sumcase TL TR T l r = Const (@{const_name "sum.sum_case"}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
+
+val App = curry op $
+
+fun mk_inj ST n i =
+ access_top_down
+ { init = (ST, I : term -> term),
+ left = (fn (T as Type ("+", [LT, RT]), inj) => (LT, inj o App (Const (@{const_name "Inl"}, LT --> T)))),
+ right =(fn (T as Type ("+", [LT, RT]), inj) => (RT, inj o App (Const (@{const_name "Inr"}, RT --> T))))} n i
+ |> snd
+
+fun mk_proj ST n i =
+ access_top_down
+ { init = (ST, I : term -> term),
+ left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name "Datatype.Projl"}, T --> LT)) o proj)),
+ right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name "Datatype.Projr"}, T --> RT)) o proj))} n i
+ |> snd
+
+fun mk_sumcases T fs =
+ BalancedTree.make (fn ((f, fT), (g, gT)) => (mk_sumcase fT gT T f g, mk_sumT fT gT))
+ (map (fn f => (f, domain_type (fastype_of f))) fs)
+ |> fst
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/termination.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,324 @@
+(* Title: HOL/Tools/Function/termination.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Context data for termination proofs
+*)
+
+
+signature TERMINATION =
+sig
+
+ type data
+ datatype cell = Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm
+
+ val mk_sumcases : data -> typ -> term list -> term
+
+ val note_measure : int -> term -> data -> data
+ val note_chain : term -> term -> thm option -> data -> data
+ val note_descent : term -> term -> term -> cell -> data -> data
+
+ val get_num_points : data -> int
+ val get_types : data -> int -> typ
+ val get_measures : data -> int -> term list
+
+ (* read from cache *)
+ val get_chain : data -> term -> term -> thm option option
+ val get_descent : data -> term -> term -> term -> cell option
+
+ (* writes *)
+ val derive_descent : theory -> tactic -> term -> term -> term -> data -> data
+ val derive_descents : theory -> tactic -> term -> data -> data
+
+ val dest_call : data -> term -> ((string * typ) list * int * term * int * term * term)
+
+ val CALLS : (term list * int -> tactic) -> int -> tactic
+
+ (* Termination tactics. Sequential composition via continuations. (2nd argument is the error continuation) *)
+ type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
+
+ val TERMINATION : Proof.context -> (data -> int -> tactic) -> int -> tactic
+
+ val REPEAT : ttac -> ttac
+
+ val wf_union_tac : Proof.context -> tactic
+end
+
+
+
+structure Termination : TERMINATION =
+struct
+
+open FundefLib
+
+val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
+structure Term2tab = TableFun(type key = term * term val ord = term2_ord);
+structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord TermOrd.fast_term_ord term2_ord);
+
+(** Analyzing binary trees **)
+
+(* Skeleton of a tree structure *)
+
+datatype skel =
+ SLeaf of int (* index *)
+| SBranch of (skel * skel)
+
+
+(* abstract make and dest functions *)
+fun mk_tree leaf branch =
+ let fun mk (SLeaf i) = leaf i
+ | mk (SBranch (s, t)) = branch (mk s, mk t)
+ in mk end
+
+
+fun dest_tree split =
+ let fun dest (SLeaf i) x = [(i, x)]
+ | dest (SBranch (s, t)) x =
+ let val (l, r) = split x
+ in dest s l @ dest t r end
+ in dest end
+
+
+(* concrete versions for sum types *)
+fun is_inj (Const ("Sum_Type.Inl", _) $ _) = true
+ | is_inj (Const ("Sum_Type.Inr", _) $ _) = true
+ | is_inj _ = false
+
+fun dest_inl (Const ("Sum_Type.Inl", _) $ t) = SOME t
+ | dest_inl _ = NONE
+
+fun dest_inr (Const ("Sum_Type.Inr", _) $ t) = SOME t
+ | dest_inr _ = NONE
+
+
+fun mk_skel ps =
+ let
+ fun skel i ps =
+ if forall is_inj ps andalso not (null ps)
+ then let
+ val (j, s) = skel i (map_filter dest_inl ps)
+ val (k, t) = skel j (map_filter dest_inr ps)
+ in (k, SBranch (s, t)) end
+ else (i + 1, SLeaf i)
+ in
+ snd (skel 0 ps)
+ end
+
+(* compute list of types for nodes *)
+fun node_types sk T = dest_tree (fn Type ("+", [LT, RT]) => (LT, RT)) sk T |> map snd
+
+(* find index and raw term *)
+fun dest_inj (SLeaf i) trm = (i, trm)
+ | dest_inj (SBranch (s, t)) trm =
+ case dest_inl trm of
+ SOME trm' => dest_inj s trm'
+ | _ => dest_inj t (the (dest_inr trm))
+
+
+
+(** Matrix cell datatype **)
+
+datatype cell = Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm;
+
+
+type data =
+ skel (* structure of the sum type encoding "program points" *)
+ * (int -> typ) (* types of program points *)
+ * (term list Inttab.table) (* measures for program points *)
+ * (thm option Term2tab.table) (* which calls form chains? *)
+ * (cell Term3tab.table) (* local descents *)
+
+
+fun map_measures f (p, T, M, C, D) = (p, T, f M, C, D)
+fun map_chains f (p, T, M, C, D) = (p, T, M, f C, D)
+fun map_descent f (p, T, M, C, D) = (p, T, M, C, f D)
+
+fun note_measure p m = map_measures (Inttab.insert_list (op aconv) (p, m))
+fun note_chain c1 c2 res = map_chains (Term2tab.update ((c1, c2), res))
+fun note_descent c m1 m2 res = map_descent (Term3tab.update ((c,(m1, m2)), res))
+
+(* Build case expression *)
+fun mk_sumcases (sk, _, _, _, _) T fs =
+ mk_tree (fn i => (nth fs i, domain_type (fastype_of (nth fs i))))
+ (fn ((f, fT), (g, gT)) => (SumTree.mk_sumcase fT gT T f g, SumTree.mk_sumT fT gT))
+ sk
+ |> fst
+
+fun mk_sum_skel rel =
+ let
+ val cs = FundefLib.dest_binop_list @{const_name Un} rel
+ fun collect_pats (Const ("Collect", _) $ Abs (_, _, c)) =
+ let
+ val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
+ = Term.strip_qnt_body "Ex" c
+ in cons r o cons l end
+ in
+ mk_skel (fold collect_pats cs [])
+ end
+
+fun create ctxt T rel =
+ let
+ val sk = mk_sum_skel rel
+ val Ts = node_types sk T
+ val M = Inttab.make (map_index (apsnd (MeasureFunctions.get_measure_functions ctxt)) Ts)
+ in
+ (sk, nth Ts, M, Term2tab.empty, Term3tab.empty)
+ end
+
+fun get_num_points (sk, _, _, _, _) =
+ let
+ fun num (SLeaf i) = i + 1
+ | num (SBranch (s, t)) = num t
+ in num sk end
+
+fun get_types (_, T, _, _, _) = T
+fun get_measures (_, _, M, _, _) = Inttab.lookup_list M
+
+fun get_chain (_, _, _, C, _) c1 c2 =
+ Term2tab.lookup C (c1, c2)
+
+fun get_descent (_, _, _, _, D) c m1 m2 =
+ Term3tab.lookup D (c, (m1, m2))
+
+fun dest_call D (Const ("Collect", _) $ Abs (_, _, c)) =
+ let
+ val n = get_num_points D
+ val (sk, _, _, _, _) = D
+ val vs = Term.strip_qnt_vars "Ex" c
+
+ (* FIXME: throw error "dest_call" for malformed terms *)
+ val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
+ = Term.strip_qnt_body "Ex" c
+ val (p, l') = dest_inj sk l
+ val (q, r') = dest_inj sk r
+ in
+ (vs, p, l', q, r', Gam)
+ end
+ | dest_call D t = error "dest_call"
+
+
+fun derive_desc_aux thy tac c (vs, p, l', q, r', Gam) m1 m2 D =
+ case get_descent D c m1 m2 of
+ SOME _ => D
+ | NONE => let
+ fun cgoal rel =
+ Term.list_all (vs,
+ Logic.mk_implies (HOLogic.mk_Trueprop Gam,
+ HOLogic.mk_Trueprop (Const (rel, @{typ "nat => nat => bool"})
+ $ (m2 $ r') $ (m1 $ l'))))
+ |> cterm_of thy
+ in
+ note_descent c m1 m2
+ (case try_proof (cgoal @{const_name HOL.less}) tac of
+ Solved thm => Less thm
+ | Stuck thm =>
+ (case try_proof (cgoal @{const_name HOL.less_eq}) tac of
+ Solved thm2 => LessEq (thm2, thm)
+ | Stuck thm2 =>
+ if prems_of thm2 = [HOLogic.Trueprop $ HOLogic.false_const]
+ then False thm2 else None (thm2, thm)
+ | _ => raise Match) (* FIXME *)
+ | _ => raise Match) D
+ end
+
+fun derive_descent thy tac c m1 m2 D =
+ derive_desc_aux thy tac c (dest_call D c) m1 m2 D
+
+(* all descents in one go *)
+fun derive_descents thy tac c D =
+ let val cdesc as (vs, p, l', q, r', Gam) = dest_call D c
+ in fold_product (derive_desc_aux thy tac c cdesc)
+ (get_measures D p) (get_measures D q) D
+ end
+
+fun CALLS tac i st =
+ if Thm.no_prems st then all_tac st
+ else case Thm.term_of (Thm.cprem_of st i) of
+ (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Un} rel, i) st
+ |_ => no_tac st
+
+type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
+
+fun TERMINATION ctxt tac =
+ SUBGOAL (fn (_ $ (Const (@{const_name "wf"}, wfT) $ rel), i) =>
+ let
+ val (T, _) = HOLogic.dest_prodT (HOLogic.dest_setT (domain_type wfT))
+ in
+ tac (create ctxt T rel) i
+ end)
+
+
+(* A tactic to convert open to closed termination goals *)
+local
+fun dest_term (t : term) = (* FIXME, cf. Lexicographic order *)
+ let
+ val (vars, prop) = FundefLib.dest_all_all t
+ val (prems, concl) = Logic.strip_horn prop
+ val (lhs, rhs) = concl
+ |> HOLogic.dest_Trueprop
+ |> HOLogic.dest_mem |> fst
+ |> HOLogic.dest_prod
+ in
+ (vars, prems, lhs, rhs)
+ end
+
+fun mk_pair_compr (T, qs, l, r, conds) =
+ let
+ val pT = HOLogic.mk_prodT (T, T)
+ val n = length qs
+ val peq = HOLogic.eq_const pT $ Bound n $ (HOLogic.pair_const T T $ l $ r)
+ val conds' = if null conds then [HOLogic.true_const] else conds
+ in
+ HOLogic.Collect_const pT $
+ Abs ("uu_", pT,
+ (foldr1 HOLogic.mk_conj (peq :: conds')
+ |> fold_rev (fn v => fn t => HOLogic.exists_const (fastype_of v) $ lambda v t) qs))
+ end
+
+in
+
+fun wf_union_tac ctxt st =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val cert = cterm_of (theory_of_thm st)
+ val ((trueprop $ (wf $ rel)) :: ineqs) = prems_of st
+
+ fun mk_compr ineq =
+ let
+ val (vars, prems, lhs, rhs) = dest_term ineq
+ in
+ mk_pair_compr (fastype_of lhs, vars, lhs, rhs, map (ObjectLogic.atomize_term thy) prems)
+ end
+
+ val relation =
+ if null ineqs then
+ Const (@{const_name Set.empty}, fastype_of rel)
+ else
+ foldr1 (HOLogic.mk_binop @{const_name Un}) (map mk_compr ineqs)
+
+ fun solve_membership_tac i =
+ (EVERY' (replicate (i - 2) (rtac @{thm UnI2})) (* pick the right component of the union *)
+ THEN' (fn j => TRY (rtac @{thm UnI1} j))
+ THEN' (rtac @{thm CollectI}) (* unfold comprehension *)
+ THEN' (fn i => REPEAT (rtac @{thm exI} i)) (* Turn existentials into schematic Vars *)
+ THEN' ((rtac @{thm refl}) (* unification instantiates all Vars *)
+ ORELSE' ((rtac @{thm conjI})
+ THEN' (rtac @{thm refl})
+ THEN' (blast_tac (local_claset_of ctxt)))) (* Solve rest of context... not very elegant *)
+ ) i
+ in
+ ((PRIMITIVE (Drule.cterm_instantiate [(cert rel, cert relation)])
+ THEN ALLGOALS (fn i => if i = 1 then all_tac else solve_membership_tac i))) st
+ end
+
+
+end
+
+
+(* continuation passing repeat combinator *)
+fun REPEAT ttac cont err_cont =
+ ttac (fn D => fn i => (REPEAT ttac cont cont D i)) err_cont
+
+
+
+
+end
--- a/src/HOL/Tools/TFL/casesplit.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/TFL/casesplit.ML Tue Jun 23 15:48:55 2009 +0100
@@ -90,7 +90,7 @@
(* get the case_thm (my version) from a type *)
fun case_thm_of_ty sgn ty =
let
- val dtypestab = DatatypePackage.get_datatypes sgn;
+ val dtypestab = Datatype.get_datatypes sgn;
val ty_str = case ty of
Type(ty_str, _) => ty_str
| TFree(s,_) => error ("Free type: " ^ s)
--- a/src/HOL/Tools/TFL/tfl.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/TFL/tfl.ML Tue Jun 23 15:48:55 2009 +0100
@@ -446,7 +446,7 @@
slow.*)
val case_ss = Simplifier.theory_context theory
(HOL_basic_ss addcongs
- (map (#weak_case_cong o snd) o Symtab.dest o DatatypePackage.get_datatypes) theory addsimps case_rewrites)
+ (map (#weak_case_cong o snd) o Symtab.dest o Datatype.get_datatypes) theory addsimps case_rewrites)
val corollaries' = map (Simplifier.simplify case_ss) corollaries
val extract = R.CONTEXT_REWRITE_RULE
(f, [R], @{thm cut_apply}, meta_tflCongs@context_congs)
--- a/src/HOL/Tools/TFL/thry.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/TFL/thry.ML Tue Jun 23 15:48:55 2009 +0100
@@ -60,20 +60,20 @@
*---------------------------------------------------------------------------*)
fun match_info thy dtco =
- case (DatatypePackage.get_datatype thy dtco,
- DatatypePackage.get_datatype_constrs thy dtco) of
+ case (Datatype.get_datatype thy dtco,
+ Datatype.get_datatype_constrs thy dtco) of
(SOME { case_name, ... }, SOME constructors) =>
SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors}
| _ => NONE;
-fun induct_info thy dtco = case DatatypePackage.get_datatype thy dtco of
+fun induct_info thy dtco = case Datatype.get_datatype thy dtco of
NONE => NONE
| SOME {nchotomy, ...} =>
SOME {nchotomy = nchotomy,
- constructors = (map Const o the o DatatypePackage.get_datatype_constrs thy) dtco};
+ constructors = (map Const o the o Datatype.get_datatype_constrs thy) dtco};
fun extract_info thy =
- let val infos = (map snd o Symtab.dest o DatatypePackage.get_datatypes) thy
+ let val infos = (map snd o Symtab.dest o Datatype.get_datatypes) thy
in {case_congs = map (mk_meta_eq o #case_cong) infos,
case_rewrites = List.concat (map (map mk_meta_eq o #case_rewrites) infos)}
end;
--- a/src/HOL/Tools/atp_manager.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/atp_manager.ML Tue Jun 23 15:48:55 2009 +0100
@@ -20,9 +20,9 @@
val info: unit -> unit
val messages: int option -> unit
type prover = int -> (thm * (string * int)) list option ->
- (int Symtab.table * bool Symtab.table) option -> string -> int ->
+ (thm * (string * int)) list option -> string -> int ->
Proof.context * (thm list * thm) ->
- bool * string * string * string vector * (int Symtab.table * bool Symtab.table)
+ bool * string * string * string vector * (thm * (string * int)) list
val add_prover: string -> prover -> theory -> theory
val print_provers: theory -> unit
val get_prover: string -> theory -> prover option
@@ -292,9 +292,9 @@
(* named provers *)
type prover = int -> (thm * (string * int)) list option ->
- (int Symtab.table * bool Symtab.table) option -> string -> int ->
+ (thm * (string * int)) list option -> string -> int ->
Proof.context * (thm list * thm) ->
- bool * string * string * string vector * (int Symtab.table * bool Symtab.table)
+ bool * string * string * string vector * (thm * (string * int)) list
fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
--- a/src/HOL/Tools/atp_minimal.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/atp_minimal.ML Tue Jun 23 15:48:55 2009 +0100
@@ -83,9 +83,9 @@
("# Cannot determine problem status within resource limit", Timeout),
("Error", Error)]
-fun produce_answer (success, message, result_string, thm_name_vec, const_counts) =
+fun produce_answer (success, message, result_string, thm_name_vec, filtered) =
if success then
- (Success (Vector.foldr op:: [] thm_name_vec, const_counts), result_string)
+ (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string)
else
let
val failure = failure_strings |> get_first (fn (s, t) =>
@@ -100,7 +100,7 @@
(* wrapper for calling external prover *)
-fun sh_test_thms prover prover_name time_limit subgoalno state counts name_thms_pairs =
+fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs =
let
val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ")
val name_thm_pairs =
@@ -109,7 +109,7 @@
val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
val (result, proof) =
produce_answer
- (prover time_limit (SOME axclauses) counts prover_name subgoalno (Proof.get_goal state))
+ (prover time_limit (SOME axclauses) filtered prover_name subgoalno (Proof.get_goal state))
val _ = println (string_of_result result)
val _ = debug proof
in
@@ -127,11 +127,12 @@
val _ = debug_fn (fn () => app (fn (n, tl) =>
(debug n; app (fn t => debug (" " ^ Display.string_of_thm t)) tl)) name_thms_pairs)
val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state
- fun test_thms counts thms = case test_thms_fun counts thms of (Success _, _) => true | _ => false
+ fun test_thms filtered thms =
+ case test_thms_fun filtered thms of (Success _, _) => true | _ => false
in
(* try prove first to check result and get used theorems *)
(case test_thms_fun NONE name_thms_pairs of
- (Success (used, const_counts), _) =>
+ (Success (used, filtered), _) =>
let
val ordered_used = order_unique used
val to_use =
@@ -139,7 +140,7 @@
filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs
else
name_thms_pairs
- val min_thms = (minimal (test_thms (SOME const_counts)) to_use)
+ val min_thms = (minimal (test_thms (SOME filtered)) to_use)
val min_names = order_unique (map fst min_thms)
val _ = println ("Minimal " ^ (length_string min_thms) ^ " theorems")
val _ = debug_fn (fn () => print_names min_thms)
--- a/src/HOL/Tools/atp_wrapper.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/atp_wrapper.ML Tue Jun 23 15:48:55 2009 +0100
@@ -41,7 +41,7 @@
(* basic template *)
fun external_prover relevance_filter preparer writer (cmd, args) find_failure produce_answer
- timeout axiom_clauses const_counts name subgoalno goal =
+ timeout axiom_clauses filtered_clauses name subgoalno goal =
let
(* path to unique problem file *)
val destdir' = ! destdir
@@ -54,37 +54,40 @@
else error ("No such directory: " ^ destdir')
end
- (* write out problem file and call prover *)
+ (* get clauses and prepare them for writing *)
val (ctxt, (chain_ths, th)) = goal
val thy = ProofContext.theory_of ctxt
val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths
- val probfile = prob_pathname subgoalno
- val fname = File.platform_path probfile
val goal_cls = #1 (ResAxioms.neg_conjecture_clauses th subgoalno)
handle THM ("assume: variables", _, _) =>
error "Sledgehammer: Goal contains type variables (TVars)"
val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) goal_cls
+ val the_filtered_clauses =
+ case filtered_clauses of
+ NONE => relevance_filter goal goal_cls
+ | SOME fcls => fcls
val the_axiom_clauses =
case axiom_clauses of
- NONE => relevance_filter goal goal_cls
+ NONE => the_filtered_clauses
| SOME axcls => axcls
- val (thm_names, clauses) = preparer goal_cls chain_ths the_axiom_clauses thy
- val the_const_counts = case const_counts of
- NONE =>
- ResHolClause.count_constants(
- case axiom_clauses of
- NONE => clauses
- | SOME axcls => #2(preparer goal_cls chain_ths (relevance_filter goal goal_cls) thy)
- )
- | SOME ccs => ccs
- val _ = writer fname the_const_counts clauses
+ val (thm_names, clauses) = preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
+
+ (* write out problem file and call prover *)
+ val probfile = prob_pathname subgoalno
+ val fname = File.platform_path probfile
+ val _ = writer fname clauses
val cmdline =
if File.exists cmd then "exec " ^ File.shell_path cmd ^ " " ^ args
else error ("Bad executable: " ^ Path.implode cmd)
val (proof, rc) = system_out (cmdline ^ " " ^ fname)
- (* remove *temporary* files *)
- val _ = if destdir' = "" then OS.FileSys.remove fname else ()
+ (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
+ val _ =
+ if destdir' = "" then OS.FileSys.remove fname
+ else
+ let val out = TextIO.openOut (fname ^ "_proof")
+ val _ = TextIO.output (out, proof)
+ in TextIO.closeOut out end
(* check for success and print out some information on failure *)
val failure = find_failure proof
@@ -94,7 +97,7 @@
else if rc <> 0 then "External prover failed: " ^ proof
else "Try this command: " ^ produce_answer name (proof, thm_names, ctxt, th, subgoalno)
val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof)
- in (success, message, proof, thm_names, the_const_counts) end;
+ in (success, message, proof, thm_names, the_filtered_clauses) end;
@@ -102,7 +105,7 @@
(* generic TPTP-based provers *)
-fun tptp_prover_opts_full max_new theory_const full command timeout ax_clauses ccs name n goal =
+fun tptp_prover_opts_full max_new theory_const full command timeout ax_clauses fcls name n goal =
external_prover
(ResAtp.get_relevant max_new theory_const)
(ResAtp.prepare_clauses false)
@@ -110,7 +113,7 @@
command
ResReconstruct.find_failure
(if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list_tstp)
- timeout ax_clauses ccs name n goal;
+ timeout ax_clauses fcls name n goal;
(*arbitrary ATP with TPTP input/output and problemfile as last argument*)
fun tptp_prover_opts max_new theory_const =
@@ -167,7 +170,7 @@
(* SPASS *)
-fun spass_opts max_new theory_const timeout ax_clauses ccs name n goal = external_prover
+fun spass_opts max_new theory_const timeout ax_clauses fcls name n goal = external_prover
(ResAtp.get_relevant max_new theory_const)
(ResAtp.prepare_clauses true)
ResHolClause.dfg_write_file
@@ -175,7 +178,7 @@
"-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout)
ResReconstruct.find_failure
ResReconstruct.lemma_list_dfg
- timeout ax_clauses ccs name n goal;
+ timeout ax_clauses fcls name n goal;
val spass = spass_opts 40 true;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/choice_specification.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,257 @@
+(* Title: HOL/Tools/choice_specification.ML
+ Author: Sebastian Skalberg, TU Muenchen
+
+Package for defining constants by specification.
+*)
+
+signature CHOICE_SPECIFICATION =
+sig
+ val add_specification: string option -> (bstring * xstring * bool) list ->
+ theory * thm -> theory * thm
+end
+
+structure Choice_Specification: CHOICE_SPECIFICATION =
+struct
+
+(* actual code *)
+
+local
+ fun mk_definitional [] arg = arg
+ | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
+ case HOLogic.dest_Trueprop (concl_of thm) of
+ Const("Ex",_) $ P =>
+ let
+ val ctype = domain_type (type_of P)
+ val cname_full = Sign.intern_const thy cname
+ val cdefname = if thname = ""
+ then Thm.def_name (Long_Name.base_name cname)
+ else thname
+ val def_eq = Logic.mk_equals (Const(cname_full,ctype),
+ HOLogic.choice_const ctype $ P)
+ val (thms, thy') = PureThy.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
+ val thm' = [thm,hd thms] MRS @{thm exE_some}
+ in
+ mk_definitional cos (thy',thm')
+ end
+ | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
+
+ fun mk_axiomatic axname cos arg =
+ let
+ fun process [] (thy,tm) =
+ let
+ val (thms, thy') = PureThy.add_axioms [((Binding.name axname, HOLogic.mk_Trueprop tm),[])] thy
+ in
+ (thy',hd thms)
+ end
+ | process ((thname,cname,covld)::cos) (thy,tm) =
+ case tm of
+ Const("Ex",_) $ P =>
+ let
+ val ctype = domain_type (type_of P)
+ val cname_full = Sign.intern_const thy cname
+ val cdefname = if thname = ""
+ then Thm.def_name (Long_Name.base_name cname)
+ else thname
+ val co = Const(cname_full,ctype)
+ val thy' = Theory.add_finals_i covld [co] thy
+ val tm' = case P of
+ Abs(_, _, bodt) => subst_bound (co, bodt)
+ | _ => P $ co
+ in
+ process cos (thy',tm')
+ end
+ | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
+ in
+ process cos arg
+ end
+
+in
+fun proc_exprop axiomatic cos arg =
+ case axiomatic of
+ SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
+ | NONE => mk_definitional cos arg
+end
+
+fun add_specification axiomatic cos arg =
+ arg |> apsnd Thm.freezeT
+ |> proc_exprop axiomatic cos
+ |> apsnd standard
+
+
+(* Collect all intances of constants in term *)
+
+fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms))
+ | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms)
+ | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
+ | collect_consts ( _,tms) = tms
+
+(* Complementing Type.varify... *)
+
+fun unvarify t fmap =
+ let
+ val fmap' = map Library.swap fmap
+ fun unthaw (f as (a, S)) =
+ (case AList.lookup (op =) fmap' a of
+ NONE => TVar f
+ | SOME (b, _) => TFree (b, S))
+ in
+ map_types (map_type_tvar unthaw) t
+ end
+
+(* The syntactic meddling needed to setup add_specification for work *)
+
+fun process_spec axiomatic cos alt_props thy =
+ let
+ fun zip3 [] [] [] = []
+ | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
+ | zip3 _ _ _ = error "Choice_Specification.process_spec internal error"
+
+ fun myfoldr f [x] = x
+ | myfoldr f (x::xs) = f (x,myfoldr f xs)
+ | myfoldr f [] = error "Choice_Specification.process_spec internal error"
+
+ val rew_imps = alt_props |>
+ map (ObjectLogic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
+ val props' = rew_imps |>
+ map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
+
+ fun proc_single prop =
+ let
+ val frees = OldTerm.term_frees prop
+ val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
+ orelse error "Specificaton: Only free variables of sort 'type' allowed"
+ val prop_closed = List.foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
+ in
+ (prop_closed,frees)
+ end
+
+ val props'' = map proc_single props'
+ val frees = map snd props''
+ val prop = myfoldr HOLogic.mk_conj (map fst props'')
+ val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
+
+ val (vmap, prop_thawed) = Type.varify [] prop
+ val thawed_prop_consts = collect_consts (prop_thawed,[])
+ val (altcos,overloaded) = Library.split_list cos
+ val (names,sconsts) = Library.split_list altcos
+ val consts = map (Syntax.read_term_global thy) sconsts
+ val _ = not (Library.exists (not o Term.is_Const) consts)
+ orelse error "Specification: Non-constant found as parameter"
+
+ fun proc_const c =
+ let
+ val (_, c') = Type.varify [] c
+ val (cname,ctyp) = dest_Const c'
+ in
+ case List.filter (fn t => let val (name,typ) = dest_Const t
+ in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
+ end) thawed_prop_consts of
+ [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
+ | [cf] => unvarify cf vmap
+ | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
+ end
+ val proc_consts = map proc_const consts
+ fun mk_exist (c,prop) =
+ let
+ val T = type_of c
+ val cname = Long_Name.base_name (fst (dest_Const c))
+ val vname = if Syntax.is_identifier cname
+ then cname
+ else "x"
+ in
+ HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
+ end
+ val ex_prop = List.foldr mk_exist prop proc_consts
+ val cnames = map (fst o dest_Const) proc_consts
+ fun post_process (arg as (thy,thm)) =
+ let
+ fun inst_all thy (thm,v) =
+ let
+ val cv = cterm_of thy v
+ val cT = ctyp_of_term cv
+ val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
+ in
+ thm RS spec'
+ end
+ fun remove_alls frees thm =
+ Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees)
+ fun process_single ((name,atts),rew_imp,frees) args =
+ let
+ fun undo_imps thm =
+ equal_elim (symmetric rew_imp) thm
+
+ fun add_final (arg as (thy, thm)) =
+ if name = ""
+ then arg |> Library.swap
+ else (writeln (" " ^ name ^ ": " ^ (Display.string_of_thm thm));
+ PureThy.store_thm (Binding.name name, thm) thy)
+ in
+ args |> apsnd (remove_alls frees)
+ |> apsnd undo_imps
+ |> apsnd standard
+ |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
+ |> add_final
+ |> Library.swap
+ end
+
+ fun process_all [proc_arg] args =
+ process_single proc_arg args
+ | process_all (proc_arg::rest) (thy,thm) =
+ let
+ val single_th = thm RS conjunct1
+ val rest_th = thm RS conjunct2
+ val (thy',_) = process_single proc_arg (thy,single_th)
+ in
+ process_all rest (thy',rest_th)
+ end
+ | process_all [] _ = error "Choice_Specification.process_spec internal error"
+ val alt_names = map fst alt_props
+ val _ = if exists (fn(name,_) => not (name = "")) alt_names
+ then writeln "specification"
+ else ()
+ in
+ arg |> apsnd Thm.freezeT
+ |> process_all (zip3 alt_names rew_imps frees)
+ end
+
+ fun after_qed [[thm]] = ProofContext.theory (fn thy =>
+ #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
+ end;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
+val opt_overloaded = P.opt_keyword "overloaded";
+
+val specification_decl =
+ P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
+ Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop)
+
+val _ =
+ OuterSyntax.command "specification" "define constants by specification" K.thy_goal
+ (specification_decl >> (fn (cos,alt_props) =>
+ Toplevel.print o (Toplevel.theory_to_proof
+ (process_spec NONE cos alt_props))))
+
+val ax_specification_decl =
+ P.name --
+ (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
+ Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop))
+
+val _ =
+ OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
+ (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
+ Toplevel.print o (Toplevel.theory_to_proof
+ (process_spec (SOME axname) cos alt_props))))
+
+end
+
+
+end
--- a/src/HOL/Tools/datatype_package/datatype_abs_proofs.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,449 +0,0 @@
-(* Title: HOL/Tools/datatype_abs_proofs.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Proofs and defintions independent of concrete representation
-of datatypes (i.e. requiring only abstract properties such as
-injectivity / distinctness of constructors and induction)
-
- - case distinction (exhaustion) theorems
- - characteristic equations for primrec combinators
- - characteristic equations for case combinators
- - equations for splitting "P (case ...)" expressions
- - "nchotomy" and "case_cong" theorems for TFL
-*)
-
-signature DATATYPE_ABS_PROOFS =
-sig
- type datatype_config = DatatypeAux.datatype_config
- type descr = DatatypeAux.descr
- type datatype_info = DatatypeAux.datatype_info
- val prove_casedist_thms : datatype_config -> string list ->
- descr list -> (string * sort) list -> thm ->
- attribute list -> theory -> thm list * theory
- val prove_primrec_thms : datatype_config -> string list ->
- descr list -> (string * sort) list ->
- datatype_info Symtab.table -> thm list list -> thm list list ->
- simpset -> thm -> theory -> (string list * thm list) * theory
- val prove_case_thms : datatype_config -> string list ->
- descr list -> (string * sort) list ->
- string list -> thm list -> theory -> (thm list list * string list) * theory
- val prove_split_thms : datatype_config -> string list ->
- descr list -> (string * sort) list ->
- thm list list -> thm list list -> thm list -> thm list list -> theory ->
- (thm * thm) list * theory
- val prove_nchotomys : datatype_config -> string list -> descr list ->
- (string * sort) list -> thm list -> theory -> thm list * theory
- val prove_weak_case_congs : string list -> descr list ->
- (string * sort) list -> theory -> thm list * theory
- val prove_case_congs : string list ->
- descr list -> (string * sort) list ->
- thm list -> thm list list -> theory -> thm list * theory
-end;
-
-structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS =
-struct
-
-open DatatypeAux;
-
-(************************ case distinction theorems ***************************)
-
-fun prove_casedist_thms (config : datatype_config) new_type_names descr sorts induct case_names_exhausts thy =
- let
- val _ = message config "Proving case distinction theorems ...";
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- val {maxidx, ...} = rep_thm induct;
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
- fun prove_casedist_thm ((i, t), T) =
- let
- val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
- Abs ("z", T', Const ("True", T''))) induct_Ps;
- val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
- Var (("P", 0), HOLogic.boolT))
- val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs)));
- val cert = cterm_of thy;
- val insts' = (map cert induct_Ps) ~~ (map cert insts);
- val induct' = refl RS ((List.nth
- (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
-
- in
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY
- [rtac induct' 1,
- REPEAT (rtac TrueI 1),
- REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
- REPEAT (rtac TrueI 1)])
- end;
-
- val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
- (DatatypeProp.make_casedists descr sorts) ~~ newTs)
- in
- thy
- |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms
- end;
-
-
-(*************************** primrec combinators ******************************)
-
-fun prove_primrec_thms (config : datatype_config) new_type_names descr sorts
- (dt_info : datatype_info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
- let
- val _ = message config "Constructing primrec combinators ...";
-
- val big_name = space_implode "_" new_type_names;
- val thy0 = add_path (#flat_names config) big_name thy;
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
-
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
- val big_rec_name' = big_name ^ "_rec_set";
- val rec_set_names' =
- if length descr' = 1 then [big_rec_name'] else
- map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
- (1 upto (length descr'));
- val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
-
- val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
-
- val rec_set_Ts = map (fn (T1, T2) =>
- reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
- val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
- (rec_set_names' ~~ rec_set_Ts);
- val rec_sets = map (fn c => list_comb (Const c, rec_fns))
- (rec_set_names ~~ rec_set_Ts);
-
- (* introduction rules for graph of primrec function *)
-
- fun make_rec_intr T rec_set ((rec_intr_ts, l), (cname, cargs)) =
- let
- fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
- let val free1 = mk_Free "x" U j
- in (case (strip_dtyp dt, strip_type U) of
- ((_, DtRec m), (Us, _)) =>
- let
- val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k;
- val i = length Us
- in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
- (map (pair "x") Us, List.nth (rec_sets', m) $
- app_bnds free1 i $ app_bnds free2 i)) :: prems,
- free1::t1s, free2::t2s)
- end
- | _ => (j + 1, k, prems, free1::t1s, t2s))
- end;
-
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val (_, _, prems, t1s, t2s) = List.foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts)
-
- in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
- (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
- list_comb (List.nth (rec_fns, l), t1s @ t2s)))], l + 1)
- end;
-
- val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) =>
- Library.foldl (make_rec_intr T set_name) (x, #3 (snd d)))
- (([], 0), descr' ~~ recTs ~~ rec_sets');
-
- val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
- (map dest_Free rec_fns)
- (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
-
- (* prove uniqueness and termination of primrec combinators *)
-
- val _ = message config "Proving termination and uniqueness of primrec functions ...";
-
- fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
- let
- val distinct_tac =
- (if i < length newTs then
- full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1
- else full_simp_tac dist_ss 1);
-
- val inject = map (fn r => r RS iffD1)
- (if i < length newTs then List.nth (constr_inject, i)
- else #inject (the (Symtab.lookup dt_info tname)));
-
- fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
- let
- val k = length (List.filter is_rec_type cargs)
-
- in (EVERY [DETERM tac,
- REPEAT (etac ex1E 1), rtac ex1I 1,
- DEPTH_SOLVE_1 (ares_tac [intr] 1),
- REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
- etac elim 1,
- REPEAT_DETERM_N j distinct_tac,
- TRY (dresolve_tac inject 1),
- REPEAT (etac conjE 1), hyp_subst_tac 1,
- REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
- TRY (hyp_subst_tac 1),
- rtac refl 1,
- REPEAT_DETERM_N (n - j - 1) distinct_tac],
- intrs, j + 1)
- end;
-
- val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs))
- ((tac, intrs, 0), constrs);
-
- in (tac', intrs') end;
-
- val rec_unique_thms =
- let
- val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
- Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
- absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
- (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
- val cert = cterm_of thy1
- val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
- ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
- val induct' = cterm_instantiate ((map cert induct_Ps) ~~
- (map cert insts)) induct;
- val (tac, _) = Library.foldl mk_unique_tac
- (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
- THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs),
- descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
-
- in split_conj_thm (SkipProof.prove_global thy1 [] []
- (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
- end;
-
- val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
-
- (* define primrec combinators *)
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.full_bname thy1)
- (if length descr' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- val (reccomb_defs, thy2) =
- thy1
- |> Sign.add_consts_i (map (fn ((name, T), T') =>
- (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
- (reccomb_names ~~ recTs ~~ rec_result_Ts))
- |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
- (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
- Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
- set $ Free ("x", T) $ Free ("y", T'))))))
- (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
- ||> parent_path (#flat_names config)
- ||> Theory.checkpoint;
-
-
- (* prove characteristic equations for primrec combinators *)
-
- val _ = message config "Proving characteristic theorems for primrec combinators ..."
-
- val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
- (fn _ => EVERY
- [rewrite_goals_tac reccomb_defs,
- rtac the1_equality 1,
- resolve_tac rec_unique_thms 1,
- resolve_tac rec_intrs 1,
- REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
- (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
-
- in
- thy2
- |> Sign.add_path (space_implode "_" new_type_names)
- |> PureThy.add_thmss [((Binding.name "recs", rec_thms),
- [Nitpick_Const_Simp_Thms.add])]
- ||> Sign.parent_path
- ||> Theory.checkpoint
- |-> (fn thms => pair (reccomb_names, Library.flat thms))
- end;
-
-
-(***************************** case combinators *******************************)
-
-fun prove_case_thms (config : datatype_config) new_type_names descr sorts reccomb_names primrec_thms thy =
- let
- val _ = message config "Proving characteristic theorems for case combinators ...";
-
- val thy1 = add_path (#flat_names config) (space_implode "_" new_type_names) thy;
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used "'t", HOLogic.typeS);
-
- fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T';
-
- val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
- in Const (@{const_name undefined}, Ts @ Ts' ---> T')
- end) constrs) descr';
-
- val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
-
- (* define case combinators via primrec combinators *)
-
- val (case_defs, thy2) = Library.foldl (fn ((defs, thy),
- ((((i, (_, _, constrs)), T), name), recname)) =>
- let
- val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
- val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
- val frees = Library.take (length cargs, frees');
- val free = mk_Free "f" (Ts ---> T') j
- in
- (free, list_abs_free (map dest_Free frees',
- list_comb (free, frees)))
- end) (constrs ~~ (1 upto length constrs)));
-
- val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
- val fns = (List.concat (Library.take (i, case_dummy_fns))) @
- fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
- val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
- val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
- val def = (Binding.name (Long_Name.base_name name ^ "_def"),
- Logic.mk_equals (list_comb (Const (name, caseT), fns1),
- list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
- fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
- val ([def_thm], thy') =
- thy
- |> Sign.declare_const [] decl |> snd
- |> (PureThy.add_defs false o map Thm.no_attributes) [def];
-
- in (defs @ [def_thm], thy')
- end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
- (Library.take (length newTs, reccomb_names)))
- ||> Theory.checkpoint;
-
- val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
- (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
- (DatatypeProp.make_cases new_type_names descr sorts thy2)
- in
- thy2
- |> Context.the_theory o fold (fold Nitpick_Const_Simp_Thms.add_thm) case_thms
- o Context.Theory
- |> parent_path (#flat_names config)
- |> store_thmss "cases" new_type_names case_thms
- |-> (fn thmss => pair (thmss, case_names))
- end;
-
-
-(******************************* case splitting *******************************)
-
-fun prove_split_thms (config : datatype_config) new_type_names descr sorts constr_inject dist_rewrites
- casedist_thms case_thms thy =
- let
- val _ = message config "Proving equations for case splitting ...";
-
- val descr' = flat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
- exhaustion), case_thms'), T) =
- let
- val cert = cterm_of thy;
- val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
- val exhaustion' = cterm_instantiate
- [(cert lhs, cert (Free ("x", T)))] exhaustion;
- val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
- (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
- in
- (SkipProof.prove_global thy [] [] t1 tacf,
- SkipProof.prove_global thy [] [] t2 tacf)
- end;
-
- val split_thm_pairs = map prove_split_thms
- ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
- dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
-
- val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
-
- in
- thy
- |> store_thms "split" new_type_names split_thms
- ||>> store_thms "split_asm" new_type_names split_asm_thms
- |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
- end;
-
-fun prove_weak_case_congs new_type_names descr sorts thy =
- let
- fun prove_weak_case_cong t =
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
-
- val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
- new_type_names descr sorts thy)
-
- in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end;
-
-(************************* additional theorems for TFL ************************)
-
-fun prove_nchotomys (config : datatype_config) new_type_names descr sorts casedist_thms thy =
- let
- val _ = message config "Proving additional theorems for TFL ...";
-
- fun prove_nchotomy (t, exhaustion) =
- let
- (* For goal i, select the correct disjunct to attack, then prove it *)
- fun tac i 0 = EVERY [TRY (rtac disjI1 i),
- hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
- | tac i n = rtac disjI2 i THEN tac i (n - 1)
- in
- SkipProof.prove_global thy [] [] t (fn _ =>
- EVERY [rtac allI 1,
- exh_tac (K exhaustion) 1,
- ALLGOALS (fn i => tac i (i-1))])
- end;
-
- val nchotomys =
- map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
-
- in thy |> store_thms "nchotomy" new_type_names nchotomys end;
-
-fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
- let
- fun prove_case_cong ((t, nchotomy), case_rewrites) =
- let
- val (Const ("==>", _) $ tm $ _) = t;
- val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
- val cert = cterm_of thy;
- val nchotomy' = nchotomy RS spec;
- val [v] = Term.add_vars (concl_of nchotomy') [];
- val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
- in
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} =>
- let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
- in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
- cut_facts_tac [nchotomy''] 1,
- REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
- REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
- end)
- end;
-
- val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
- new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
-
- in thy |> store_thms "case_cong" new_type_names case_congs end;
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_aux.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,375 +0,0 @@
-(* Title: HOL/Tools/datatype_aux.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Auxiliary functions for defining datatypes.
-*)
-
-signature DATATYPE_AUX =
-sig
- type datatype_config
- val default_datatype_config : datatype_config
- val message : datatype_config -> string -> unit
-
- val add_path : bool -> string -> theory -> theory
- val parent_path : bool -> theory -> theory
-
- val store_thmss_atts : string -> string list -> attribute list list -> thm list list
- -> theory -> thm list list * theory
- val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
- val store_thms_atts : string -> string list -> attribute list list -> thm list
- -> theory -> thm list * theory
- val store_thms : string -> string list -> thm list -> theory -> thm list * theory
-
- val split_conj_thm : thm -> thm list
- val mk_conj : term list -> term
- val mk_disj : term list -> term
-
- val app_bnds : term -> int -> term
-
- val cong_tac : int -> tactic
- val indtac : thm -> string list -> int -> tactic
- val exh_tac : (string -> thm) -> int -> tactic
-
- datatype simproc_dist = FewConstrs of thm list
- | ManyConstrs of thm * simpset;
-
- datatype dtyp =
- DtTFree of string
- | DtType of string * (dtyp list)
- | DtRec of int;
- type descr
- type datatype_info
-
- exception Datatype
- exception Datatype_Empty of string
- val name_of_typ : typ -> string
- val dtyp_of_typ : (string * string list) list -> typ -> dtyp
- val mk_Free : string -> typ -> int -> term
- val is_rec_type : dtyp -> bool
- val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
- val dest_DtTFree : dtyp -> string
- val dest_DtRec : dtyp -> int
- val strip_dtyp : dtyp -> dtyp list * dtyp
- val body_index : dtyp -> int
- val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
- val get_nonrec_types : descr -> (string * sort) list -> typ list
- val get_branching_types : descr -> (string * sort) list -> typ list
- val get_arities : descr -> int list
- val get_rec_types : descr -> (string * sort) list -> typ list
- val interpret_construction : descr -> (string * sort) list
- -> { atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a }
- -> ((string * Term.typ list) * (string * 'a list) list) list
- val check_nonempty : descr list -> unit
- val unfold_datatypes :
- theory -> descr -> (string * sort) list -> datatype_info Symtab.table ->
- descr -> int -> descr list * int
-end;
-
-structure DatatypeAux : DATATYPE_AUX =
-struct
-
-(* datatype option flags *)
-
-type datatype_config = { strict: bool, flat_names: bool, quiet: bool };
-val default_datatype_config : datatype_config =
- { strict = true, flat_names = false, quiet = false };
-fun message ({ quiet, ...} : datatype_config) s =
- if quiet then () else writeln s;
-
-fun add_path flat_names s = if flat_names then I else Sign.add_path s;
-fun parent_path flat_names = if flat_names then I else Sign.parent_path;
-
-
-(* store theorems in theory *)
-
-fun store_thmss_atts label tnames attss thmss =
- fold_map (fn ((tname, atts), thms) =>
- Sign.add_path tname
- #> PureThy.add_thmss [((Binding.name label, thms), atts)]
- #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
- ##> Theory.checkpoint;
-
-fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
-
-fun store_thms_atts label tnames attss thmss =
- fold_map (fn ((tname, atts), thms) =>
- Sign.add_path tname
- #> PureThy.add_thms [((Binding.name label, thms), atts)]
- #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
- ##> Theory.checkpoint;
-
-fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
-
-
-(* split theorem thm_1 & ... & thm_n into n theorems *)
-
-fun split_conj_thm th =
- ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
-
-val mk_conj = foldr1 (HOLogic.mk_binop "op &");
-val mk_disj = foldr1 (HOLogic.mk_binop "op |");
-
-fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
-
-
-fun cong_tac i st = (case Logic.strip_assums_concl
- (List.nth (prems_of st, i - 1)) of
- _ $ (_ $ (f $ x) $ (g $ y)) =>
- let
- val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
- val _ $ (_ $ (f' $ x') $ (g' $ y')) =
- Logic.strip_assums_concl (prop_of cong');
- val insts = map (pairself (cterm_of (Thm.theory_of_thm st)) o
- apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
- apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
- in compose_tac (false, cterm_instantiate insts cong', 2) i st
- handle THM _ => no_tac st
- end
- | _ => no_tac st);
-
-(* instantiate induction rule *)
-
-fun indtac indrule indnames i st =
- let
- val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
- val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
- (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
- val getP = if can HOLogic.dest_imp (hd ts) then
- (apfst SOME) o HOLogic.dest_imp else pair NONE;
- val flt = if null indnames then I else
- filter (fn Free (s, _) => s mem indnames | _ => false);
- fun abstr (t1, t2) = (case t1 of
- NONE => (case flt (OldTerm.term_frees t2) of
- [Free (s, T)] => SOME (absfree (s, T, t2))
- | _ => NONE)
- | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
- val cert = cterm_of (Thm.theory_of_thm st);
- val insts = List.mapPartial (fn (t, u) => case abstr (getP u) of
- NONE => NONE
- | SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u')) (ts ~~ ts');
- val indrule' = cterm_instantiate insts indrule
- in
- rtac indrule' i st
- end;
-
-(* perform exhaustive case analysis on last parameter of subgoal i *)
-
-fun exh_tac exh_thm_of i state =
- let
- val thy = Thm.theory_of_thm state;
- val prem = nth (prems_of state) (i - 1);
- val params = Logic.strip_params prem;
- val (_, Type (tname, _)) = hd (rev params);
- val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
- val prem' = hd (prems_of exhaustion);
- val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
- val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs),
- cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t))
- (Bound 0) params))] exhaustion
- in compose_tac (false, exhaustion', nprems_of exhaustion) i state
- end;
-
-(* handling of distinctness theorems *)
-
-datatype simproc_dist = FewConstrs of thm list
- | ManyConstrs of thm * simpset;
-
-(********************** Internal description of datatypes *********************)
-
-datatype dtyp =
- DtTFree of string
- | DtType of string * (dtyp list)
- | DtRec of int;
-
-(* information about datatypes *)
-
-(* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
-type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
-
-type datatype_info =
- {index : int,
- alt_names : string list option,
- descr : descr,
- sorts : (string * sort) list,
- rec_names : string list,
- rec_rewrites : thm list,
- case_name : string,
- case_rewrites : thm list,
- induction : thm,
- exhaustion : thm,
- distinct : simproc_dist,
- inject : thm list,
- nchotomy : thm,
- case_cong : thm,
- weak_case_cong : thm};
-
-fun mk_Free s T i = Free (s ^ (string_of_int i), T);
-
-fun subst_DtTFree _ substs (T as (DtTFree name)) =
- AList.lookup (op =) substs name |> the_default T
- | subst_DtTFree i substs (DtType (name, ts)) =
- DtType (name, map (subst_DtTFree i substs) ts)
- | subst_DtTFree i _ (DtRec j) = DtRec (i + j);
-
-exception Datatype;
-exception Datatype_Empty of string;
-
-fun dest_DtTFree (DtTFree a) = a
- | dest_DtTFree _ = raise Datatype;
-
-fun dest_DtRec (DtRec i) = i
- | dest_DtRec _ = raise Datatype;
-
-fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
- | is_rec_type (DtRec _) = true
- | is_rec_type _ = false;
-
-fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
- | strip_dtyp T = ([], T);
-
-val body_index = dest_DtRec o snd o strip_dtyp;
-
-fun mk_fun_dtyp [] U = U
- | mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
-
-fun name_of_typ (Type (s, Ts)) =
- let val s' = Long_Name.base_name s
- in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
- [if Syntax.is_identifier s' then s' else "x"])
- end
- | name_of_typ _ = "";
-
-fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
- | dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
- | dtyp_of_typ new_dts (Type (tname, Ts)) =
- (case AList.lookup (op =) new_dts tname of
- NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
- | SOME vs => if map (try (fst o dest_TFree)) Ts = map SOME vs then
- DtRec (find_index (curry op = tname o fst) new_dts)
- else error ("Illegal occurrence of recursive type " ^ tname));
-
-fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
- | typ_of_dtyp descr sorts (DtRec i) =
- let val (s, ds, _) = (the o AList.lookup (op =) descr) i
- in Type (s, map (typ_of_dtyp descr sorts) ds) end
- | typ_of_dtyp descr sorts (DtType (s, ds)) =
- Type (s, map (typ_of_dtyp descr sorts) ds);
-
-(* find all non-recursive types in datatype description *)
-
-fun get_nonrec_types descr sorts =
- map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
- Library.foldl (fn (Ts', (_, cargs)) =>
- filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
-
-(* get all recursive types in datatype description *)
-
-fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
- Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
-
-(* get all branching types *)
-
-fun get_branching_types descr sorts =
- map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
- fold (fn (_, cargs) => fold (strip_dtyp #> fst #> fold (insert op =)) cargs)
- constrs) descr []);
-
-fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
- fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
- (List.filter is_rec_type cargs))) constrs) descr [];
-
-(* interpret construction of datatype *)
-
-fun interpret_construction descr vs { atyp, dtyp } =
- let
- val typ_of_dtyp = typ_of_dtyp descr vs;
- fun interpT dT = case strip_dtyp dT
- of (dTs, DtRec l) =>
- let
- val (tyco, dTs', _) = (the o AList.lookup (op =) descr) l;
- val Ts = map typ_of_dtyp dTs;
- val Ts' = map typ_of_dtyp dTs';
- val is_proper = forall (can dest_TFree) Ts';
- in dtyp Ts (l, is_proper) (tyco, Ts') end
- | _ => atyp (typ_of_dtyp dT);
- fun interpC (c, dTs) = (c, map interpT dTs);
- fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
- in map interpD descr end;
-
-(* nonemptiness check for datatypes *)
-
-fun check_nonempty descr =
- let
- val descr' = List.concat descr;
- fun is_nonempty_dt is i =
- let
- val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
- fun arg_nonempty (_, DtRec i) = if i mem is then false
- else is_nonempty_dt (i::is) i
- | arg_nonempty _ = true;
- in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
- end
- in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
- (fn (_, (s, _, _)) => raise Datatype_Empty s)
- end;
-
-(* unfold a list of mutually recursive datatype specifications *)
-(* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
-(* need to be unfolded *)
-
-fun unfold_datatypes sign orig_descr sorts (dt_info : datatype_info Symtab.table) descr i =
- let
- fun typ_error T msg = error ("Non-admissible type expression\n" ^
- Syntax.string_of_typ_global sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
-
- fun get_dt_descr T i tname dts =
- (case Symtab.lookup dt_info tname of
- NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
- \ nested recursion")
- | (SOME {index, descr, ...}) =>
- let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
- val subst = ((map dest_DtTFree vars) ~~ dts) handle Library.UnequalLengths =>
- typ_error T ("Type constructor " ^ tname ^ " used with wrong\
- \ number of arguments")
- in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
- (tn, map (subst_DtTFree i subst) args,
- map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
- end);
-
- (* unfold a single constructor argument *)
-
- fun unfold_arg ((i, Ts, descrs), T) =
- if is_rec_type T then
- let val (Us, U) = strip_dtyp T
- in if exists is_rec_type Us then
- typ_error T "Non-strictly positive recursive occurrence of type"
- else (case U of
- DtType (tname, dts) =>
- let
- val (index, descr) = get_dt_descr T i tname dts;
- val (descr', i') = unfold_datatypes sign orig_descr sorts
- dt_info descr (i + length descr)
- in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
- | _ => (i, Ts @ [T], descrs))
- end
- else (i, Ts @ [T], descrs);
-
- (* unfold a constructor *)
-
- fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
- let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
- in (i', constrs @ [(cname, cargs')], descrs') end;
-
- (* unfold a single datatype *)
-
- fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
- let val (i', constrs', descrs') =
- Library.foldl unfold_constr ((i, [], descrs), constrs)
- in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
- end;
-
- val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
-
- in (descr' :: descrs, i') end;
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_case.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,469 +0,0 @@
-(* Title: HOL/Tools/datatype_case.ML
- Author: Konrad Slind, Cambridge University Computer Laboratory
- Author: Stefan Berghofer, TU Muenchen
-
-Nested case expressions on datatypes.
-*)
-
-signature DATATYPE_CASE =
-sig
- val make_case: (string -> DatatypeAux.datatype_info option) ->
- Proof.context -> bool -> string list -> term -> (term * term) list ->
- term * (term * (int * bool)) list
- val dest_case: (string -> DatatypeAux.datatype_info option) -> bool ->
- string list -> term -> (term * (term * term) list) option
- val strip_case: (string -> DatatypeAux.datatype_info option) -> bool ->
- term -> (term * (term * term) list) option
- val case_tr: bool -> (theory -> string -> DatatypeAux.datatype_info option)
- -> Proof.context -> term list -> term
- val case_tr': (theory -> string -> DatatypeAux.datatype_info option) ->
- string -> Proof.context -> term list -> term
-end;
-
-structure DatatypeCase : DATATYPE_CASE =
-struct
-
-exception CASE_ERROR of string * int;
-
-fun match_type thy pat ob = Sign.typ_match thy (pat, ob) Vartab.empty;
-
-(*---------------------------------------------------------------------------
- * Get information about datatypes
- *---------------------------------------------------------------------------*)
-
-fun ty_info (tab : string -> DatatypeAux.datatype_info option) s =
- case tab s of
- SOME {descr, case_name, index, sorts, ...} =>
- let
- val (_, (tname, dts, constrs)) = nth descr index;
- val mk_ty = DatatypeAux.typ_of_dtyp descr sorts;
- val T = Type (tname, map mk_ty dts)
- in
- SOME {case_name = case_name,
- constructors = map (fn (cname, dts') =>
- Const (cname, Logic.varifyT (map mk_ty dts' ---> T))) constrs}
- end
- | NONE => NONE;
-
-
-(*---------------------------------------------------------------------------
- * Each pattern carries with it a tag (i,b) where
- * i is the clause it came from and
- * b=true indicates that clause was given by the user
- * (or is an instantiation of a user supplied pattern)
- * b=false --> i = ~1
- *---------------------------------------------------------------------------*)
-
-fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
-
-fun row_of_pat x = fst (snd x);
-
-fun add_row_used ((prfx, pats), (tm, tag)) =
- fold Term.add_free_names (tm :: pats @ prfx);
-
-(* try to preserve names given by user *)
-fun default_names names ts =
- map (fn ("", Free (name', _)) => name' | (name, _) => name) (names ~~ ts);
-
-fun strip_constraints (Const ("_constrain", _) $ t $ tT) =
- strip_constraints t ||> cons tT
- | strip_constraints t = (t, []);
-
-fun mk_fun_constrain tT t = Syntax.const "_constrain" $ t $
- (Syntax.free "fun" $ tT $ Syntax.free "dummy");
-
-
-(*---------------------------------------------------------------------------
- * Produce an instance of a constructor, plus genvars for its arguments.
- *---------------------------------------------------------------------------*)
-fun fresh_constr ty_match ty_inst colty used c =
- let
- val (_, Ty) = dest_Const c
- val Ts = binder_types Ty;
- val names = Name.variant_list used
- (DatatypeProp.make_tnames (map Logic.unvarifyT Ts));
- val ty = body_type Ty;
- val ty_theta = ty_match ty colty handle Type.TYPE_MATCH =>
- raise CASE_ERROR ("type mismatch", ~1)
- val c' = ty_inst ty_theta c
- val gvars = map (ty_inst ty_theta o Free) (names ~~ Ts)
- in (c', gvars)
- end;
-
-
-(*---------------------------------------------------------------------------
- * Goes through a list of rows and picks out the ones beginning with a
- * pattern with constructor = name.
- *---------------------------------------------------------------------------*)
-fun mk_group (name, T) rows =
- let val k = length (binder_types T)
- in fold (fn (row as ((prfx, p :: rst), rhs as (_, (i, _)))) =>
- fn ((in_group, not_in_group), (names, cnstrts)) => (case strip_comb p of
- (Const (name', _), args) =>
- if name = name' then
- if length args = k then
- let val (args', cnstrts') = split_list (map strip_constraints args)
- in
- ((((prfx, args' @ rst), rhs) :: in_group, not_in_group),
- (default_names names args', map2 append cnstrts cnstrts'))
- end
- else raise CASE_ERROR
- ("Wrong number of arguments for constructor " ^ name, i)
- else ((in_group, row :: not_in_group), (names, cnstrts))
- | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
- rows (([], []), (replicate k "", replicate k [])) |>> pairself rev
- end;
-
-(*---------------------------------------------------------------------------
- * Partition the rows. Not efficient: we should use hashing.
- *---------------------------------------------------------------------------*)
-fun partition _ _ _ _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
- | partition ty_match ty_inst type_of used constructors colty res_ty
- (rows as (((prfx, _ :: rstp), _) :: _)) =
- let
- fun part {constrs = [], rows = [], A} = rev A
- | part {constrs = [], rows = (_, (_, (i, _))) :: _, A} =
- raise CASE_ERROR ("Not a constructor pattern", i)
- | part {constrs = c :: crst, rows, A} =
- let
- val ((in_group, not_in_group), (names, cnstrts)) =
- mk_group (dest_Const c) rows;
- val used' = fold add_row_used in_group used;
- val (c', gvars) = fresh_constr ty_match ty_inst colty used' c;
- val in_group' =
- if null in_group (* Constructor not given *)
- then
- let
- val Ts = map type_of rstp;
- val xs = Name.variant_list
- (fold Term.add_free_names gvars used')
- (replicate (length rstp) "x")
- in
- [((prfx, gvars @ map Free (xs ~~ Ts)),
- (Const ("HOL.undefined", res_ty), (~1, false)))]
- end
- else in_group
- in
- part{constrs = crst,
- rows = not_in_group,
- A = {constructor = c',
- new_formals = gvars,
- names = names,
- constraints = cnstrts,
- group = in_group'} :: A}
- end
- in part {constrs = constructors, rows = rows, A = []}
- end;
-
-(*---------------------------------------------------------------------------
- * Misc. routines used in mk_case
- *---------------------------------------------------------------------------*)
-
-fun mk_pat ((c, c'), l) =
- let
- val L = length (binder_types (fastype_of c))
- fun build (prfx, tag, plist) =
- let val (args, plist') = chop L plist
- in (prfx, tag, list_comb (c', args) :: plist') end
- in map build l end;
-
-fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
- | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
-
-fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
- | v_to_pats _ = raise CASE_ERROR ("mk_case: v_to_pats", ~1);
-
-
-(*----------------------------------------------------------------------------
- * Translation of pattern terms into nested case expressions.
- *
- * This performs the translation and also builds the full set of patterns.
- * Thus it supports the construction of induction theorems even when an
- * incomplete set of patterns is given.
- *---------------------------------------------------------------------------*)
-
-fun mk_case tab ctxt ty_match ty_inst type_of used range_ty =
- let
- val name = Name.variant used "a";
- fun expand constructors used ty ((_, []), _) =
- raise CASE_ERROR ("mk_case: expand_var_row", ~1)
- | expand constructors used ty (row as ((prfx, p :: rst), rhs)) =
- if is_Free p then
- let
- val used' = add_row_used row used;
- fun expnd c =
- let val capp =
- list_comb (fresh_constr ty_match ty_inst ty used' c)
- in ((prfx, capp :: rst), pattern_subst [(p, capp)] rhs)
- end
- in map expnd constructors end
- else [row]
- fun mk {rows = [], ...} = raise CASE_ERROR ("no rows", ~1)
- | mk {path = [], rows = ((prfx, []), (tm, tag)) :: _} = (* Done *)
- ([(prfx, tag, [])], tm)
- | mk {path, rows as ((row as ((_, [Free _]), _)) :: _ :: _)} =
- mk {path = path, rows = [row]}
- | mk {path = u :: rstp, rows as ((_, _ :: _), _) :: _} =
- let val col0 = map (fn ((_, p :: _), (_, (i, _))) => (p, i)) rows
- in case Option.map (apfst head_of)
- (find_first (not o is_Free o fst) col0) of
- NONE =>
- let
- val rows' = map (fn ((v, _), row) => row ||>
- pattern_subst [(v, u)] |>> v_to_prfx) (col0 ~~ rows);
- val (pref_patl, tm) = mk {path = rstp, rows = rows'}
- in (map v_to_pats pref_patl, tm) end
- | SOME (Const (cname, cT), i) => (case ty_info tab cname of
- NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ cname, i)
- | SOME {case_name, constructors} =>
- let
- val pty = body_type cT;
- val used' = fold Term.add_free_names rstp used;
- val nrows = maps (expand constructors used' pty) rows;
- val subproblems = partition ty_match ty_inst type_of used'
- constructors pty range_ty nrows;
- val new_formals = map #new_formals subproblems
- val constructors' = map #constructor subproblems
- val news = map (fn {new_formals, group, ...} =>
- {path = new_formals @ rstp, rows = group}) subproblems;
- val (pat_rect, dtrees) = split_list (map mk news);
- val case_functions = map2
- (fn {new_formals, names, constraints, ...} =>
- fold_rev (fn ((x as Free (_, T), s), cnstrts) => fn t =>
- Abs (if s = "" then name else s, T,
- abstract_over (x, t)) |>
- fold mk_fun_constrain cnstrts)
- (new_formals ~~ names ~~ constraints))
- subproblems dtrees;
- val types = map type_of (case_functions @ [u]);
- val case_const = Const (case_name, types ---> range_ty)
- val tree = list_comb (case_const, case_functions @ [u])
- val pat_rect1 = flat (map mk_pat
- (constructors ~~ constructors' ~~ pat_rect))
- in (pat_rect1, tree)
- end)
- | SOME (t, i) => raise CASE_ERROR ("Not a datatype constructor: " ^
- Syntax.string_of_term ctxt t, i)
- end
- | mk _ = raise CASE_ERROR ("Malformed row matrix", ~1)
- in mk
- end;
-
-fun case_error s = error ("Error in case expression:\n" ^ s);
-
-(* Repeated variable occurrences in a pattern are not allowed. *)
-fun no_repeat_vars ctxt pat = fold_aterms
- (fn x as Free (s, _) => (fn xs =>
- if member op aconv xs x then
- case_error (quote s ^ " occurs repeatedly in the pattern " ^
- quote (Syntax.string_of_term ctxt pat))
- else x :: xs)
- | _ => I) pat [];
-
-fun gen_make_case ty_match ty_inst type_of tab ctxt err used x clauses =
- let
- fun string_of_clause (pat, rhs) = Syntax.string_of_term ctxt
- (Syntax.const "_case1" $ pat $ rhs);
- val _ = map (no_repeat_vars ctxt o fst) clauses;
- val rows = map_index (fn (i, (pat, rhs)) =>
- (([], [pat]), (rhs, (i, true)))) clauses;
- val rangeT = (case distinct op = (map (type_of o snd) clauses) of
- [] => case_error "no clauses given"
- | [T] => T
- | _ => case_error "all cases must have the same result type");
- val used' = fold add_row_used rows used;
- val (patts, case_tm) = mk_case tab ctxt ty_match ty_inst type_of
- used' rangeT {path = [x], rows = rows}
- handle CASE_ERROR (msg, i) => case_error (msg ^
- (if i < 0 then ""
- else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
- val patts1 = map
- (fn (_, tag, [pat]) => (pat, tag)
- | _ => case_error "error in pattern-match translation") patts;
- val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1
- val finals = map row_of_pat patts2
- val originals = map (row_of_pat o #2) rows
- val _ = case originals \\ finals of
- [] => ()
- | is => (if err then case_error else warning)
- ("The following clauses are redundant (covered by preceding clauses):\n" ^
- cat_lines (map (string_of_clause o nth clauses) is));
- in
- (case_tm, patts2)
- end;
-
-fun make_case tab ctxt = gen_make_case
- (match_type (ProofContext.theory_of ctxt)) Envir.subst_TVars fastype_of tab ctxt;
-val make_case_untyped = gen_make_case (K (K Vartab.empty))
- (K (Term.map_types (K dummyT))) (K dummyT);
-
-
-(* parse translation *)
-
-fun case_tr err tab_of ctxt [t, u] =
- let
- val thy = ProofContext.theory_of ctxt;
- (* replace occurrences of dummy_pattern by distinct variables *)
- (* internalize constant names *)
- fun prep_pat ((c as Const ("_constrain", _)) $ t $ tT) used =
- let val (t', used') = prep_pat t used
- in (c $ t' $ tT, used') end
- | prep_pat (Const ("dummy_pattern", T)) used =
- let val x = Name.variant used "x"
- in (Free (x, T), x :: used) end
- | prep_pat (Const (s, T)) used =
- (case try (unprefix Syntax.constN) s of
- SOME c => (Const (c, T), used)
- | NONE => (Const (Sign.intern_const thy s, T), used))
- | prep_pat (v as Free (s, T)) used =
- let val s' = Sign.intern_const thy s
- in
- if Sign.declared_const thy s' then
- (Const (s', T), used)
- else (v, used)
- end
- | prep_pat (t $ u) used =
- let
- val (t', used') = prep_pat t used;
- val (u', used'') = prep_pat u used'
- in
- (t' $ u', used'')
- end
- | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
- fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
- let val (l', cnstrts) = strip_constraints l
- in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
- end
- | dest_case1 t = case_error "dest_case1";
- fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
- | dest_case2 t = [t];
- val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
- val (case_tm, _) = make_case_untyped (tab_of thy) ctxt err []
- (fold (fn tT => fn t => Syntax.const "_constrain" $ t $ tT)
- (flat cnstrts) t) cases;
- in case_tm end
- | case_tr _ _ _ ts = case_error "case_tr";
-
-
-(*---------------------------------------------------------------------------
- * Pretty printing of nested case expressions
- *---------------------------------------------------------------------------*)
-
-(* destruct one level of pattern matching *)
-
-fun gen_dest_case name_of type_of tab d used t =
- case apfst name_of (strip_comb t) of
- (SOME cname, ts as _ :: _) =>
- let
- val (fs, x) = split_last ts;
- fun strip_abs i t =
- let
- val zs = strip_abs_vars t;
- val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
- val (xs, ys) = chop i zs;
- val u = list_abs (ys, strip_abs_body t);
- val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
- (map fst xs) ~~ map snd xs)
- in (xs', subst_bounds (rev xs', u)) end;
- fun is_dependent i t =
- let val k = length (strip_abs_vars t) - i
- in k < 0 orelse exists (fn j => j >= k)
- (loose_bnos (strip_abs_body t))
- end;
- fun count_cases (_, _, true) = I
- | count_cases (c, (_, body), false) =
- AList.map_default op aconv (body, []) (cons c);
- val is_undefined = name_of #> equal (SOME "HOL.undefined");
- fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
- in case ty_info tab cname of
- SOME {constructors, case_name} =>
- if length fs = length constructors then
- let
- val cases = map (fn (Const (s, U), t) =>
- let
- val k = length (binder_types U);
- val p as (xs, _) = strip_abs k t
- in
- (Const (s, map type_of xs ---> type_of x),
- p, is_dependent k t)
- end) (constructors ~~ fs);
- val cases' = sort (int_ord o swap o pairself (length o snd))
- (fold_rev count_cases cases []);
- val R = type_of t;
- val dummy = if d then Const ("dummy_pattern", R)
- else Free (Name.variant used "x", R)
- in
- SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
- SOME (_, cs) =>
- if length cs = length constructors then [hd cases]
- else filter_out (fn (_, (_, body), _) => is_undefined body) cases
- | NONE => case cases' of
- [] => cases
- | (default, cs) :: _ =>
- if length cs = 1 then cases
- else if length cs = length constructors then
- [hd cases, (dummy, ([], default), false)]
- else
- filter_out (fn (c, _, _) => member op aconv cs c) cases @
- [(dummy, ([], default), false)]))
- end handle CASE_ERROR _ => NONE
- else NONE
- | _ => NONE
- end
- | _ => NONE;
-
-val dest_case = gen_dest_case (try (dest_Const #> fst)) fastype_of;
-val dest_case' = gen_dest_case
- (try (dest_Const #> fst #> unprefix Syntax.constN)) (K dummyT);
-
-
-(* destruct nested patterns *)
-
-fun strip_case'' dest (pat, rhs) =
- case dest (Term.add_free_names pat []) rhs of
- SOME (exp as Free _, clauses) =>
- if member op aconv (OldTerm.term_frees pat) exp andalso
- not (exists (fn (_, rhs') =>
- member op aconv (OldTerm.term_frees rhs') exp) clauses)
- then
- maps (strip_case'' dest) (map (fn (pat', rhs') =>
- (subst_free [(exp, pat')] pat, rhs')) clauses)
- else [(pat, rhs)]
- | _ => [(pat, rhs)];
-
-fun gen_strip_case dest t = case dest [] t of
- SOME (x, clauses) =>
- SOME (x, maps (strip_case'' dest) clauses)
- | NONE => NONE;
-
-val strip_case = gen_strip_case oo dest_case;
-val strip_case' = gen_strip_case oo dest_case';
-
-
-(* print translation *)
-
-fun case_tr' tab_of cname ctxt ts =
- let
- val thy = ProofContext.theory_of ctxt;
- val consts = ProofContext.consts_of ctxt;
- fun mk_clause (pat, rhs) =
- let val xs = Term.add_frees pat []
- in
- Syntax.const "_case1" $
- map_aterms
- (fn Free p => Syntax.mark_boundT p
- | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
- | t => t) pat $
- map_aterms
- (fn x as Free (s, T) =>
- if member (op =) xs (s, T) then Syntax.mark_bound s else x
- | t => t) rhs
- end
- in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
- SOME (x, clauses) => Syntax.const "_case_syntax" $ x $
- foldr1 (fn (t, u) => Syntax.const "_case2" $ t $ u)
- (map mk_clause clauses)
- | NONE => raise Match
- end;
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_codegen.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,455 +0,0 @@
-(* Title: HOL/Tools/datatype_codegen.ML
- Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
-
-Code generator facilities for inductive datatypes.
-*)
-
-signature DATATYPE_CODEGEN =
-sig
- val find_shortest_path: DatatypeAux.descr -> int -> (string * int) option
- val mk_eq_eqns: theory -> string -> (thm * bool) list
- val mk_case_cert: theory -> string -> thm
- val setup: theory -> theory
-end;
-
-structure DatatypeCodegen : DATATYPE_CODEGEN =
-struct
-
-(** find shortest path to constructor with no recursive arguments **)
-
-fun find_nonempty (descr: DatatypeAux.descr) is i =
- let
- val (_, _, constrs) = the (AList.lookup (op =) descr i);
- fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i
- then NONE
- else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i)
- | arg_nonempty _ = SOME 0;
- fun max xs = Library.foldl
- (fn (NONE, _) => NONE
- | (SOME i, SOME j) => SOME (Int.max (i, j))
- | (_, NONE) => NONE) (SOME 0, xs);
- val xs = sort (int_ord o pairself snd)
- (map_filter (fn (s, dts) => Option.map (pair s)
- (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
- in case xs of [] => NONE | x :: _ => SOME x end;
-
-fun find_shortest_path descr i = find_nonempty descr [i] i;
-
-
-(** SML code generator **)
-
-open Codegen;
-
-(* datatype definition *)
-
-fun add_dt_defs thy defs dep module (descr: DatatypeAux.descr) sorts gr =
- let
- val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
- val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
- exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
-
- val (_, (tname, _, _)) :: _ = descr';
- val node_id = tname ^ " (type)";
- val module' = if_library (thyname_of_type thy tname) module;
-
- fun mk_dtdef prfx [] gr = ([], gr)
- | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr =
- let
- val tvs = map DatatypeAux.dest_DtTFree dts;
- val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
- val ((_, type_id), gr') = mk_type_id module' tname gr;
- val (ps, gr'') = gr' |>
- fold_map (fn (cname, cargs) =>
- fold_map (invoke_tycodegen thy defs node_id module' false)
- cargs ##>>
- mk_const_id module' cname) cs';
- val (rest, gr''') = mk_dtdef "and " xs gr''
- in
- (Pretty.block (str prfx ::
- (if null tvs then [] else
- [mk_tuple (map str tvs), str " "]) @
- [str (type_id ^ " ="), Pretty.brk 1] @
- List.concat (separate [Pretty.brk 1, str "| "]
- (map (fn (ps', (_, cname)) => [Pretty.block
- (str cname ::
- (if null ps' then [] else
- List.concat ([str " of", Pretty.brk 1] ::
- separate [str " *", Pretty.brk 1]
- (map single ps'))))]) ps))) :: rest, gr''')
- end;
-
- fun mk_constr_term cname Ts T ps =
- List.concat (separate [str " $", Pretty.brk 1]
- ([str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
- mk_type false (Ts ---> T), str ")"] :: ps));
-
- fun mk_term_of_def gr prfx [] = []
- | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) =
- let
- val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
- val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val T = Type (tname, dts');
- val rest = mk_term_of_def gr "and " xs;
- val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx =>
- let val args = map (fn i =>
- str ("x" ^ string_of_int i)) (1 upto length Ts)
- in (Pretty.blk (4,
- [str prfx, mk_term_of gr module' false T, Pretty.brk 1,
- if null Ts then str (snd (get_const_id gr cname))
- else parens (Pretty.block
- [str (snd (get_const_id gr cname)),
- Pretty.brk 1, mk_tuple args]),
- str " =", Pretty.brk 1] @
- mk_constr_term cname Ts T
- (map2 (fn x => fn U => [Pretty.block [mk_term_of gr module' false U,
- Pretty.brk 1, x]]) args Ts)), " | ")
- end) cs' prfx
- in eqs @ rest end;
-
- fun mk_gen_of_def gr prfx [] = []
- | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
- let
- val tvs = map DatatypeAux.dest_DtTFree dts;
- val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val T = Type (tname, Us);
- val (cs1, cs2) =
- List.partition (exists DatatypeAux.is_rec_type o snd) cs;
- val SOME (cname, _) = find_shortest_path descr i;
-
- fun mk_delay p = Pretty.block
- [str "fn () =>", Pretty.brk 1, p];
-
- fun mk_force p = Pretty.block [p, Pretty.brk 1, str "()"];
-
- fun mk_constr s b (cname, dts) =
- let
- val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s
- (DatatypeAux.typ_of_dtyp descr sorts dt))
- [str (if b andalso DatatypeAux.is_rec_type dt then "0"
- else "j")]) dts;
- val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val xs = map str
- (DatatypeProp.indexify_names (replicate (length dts) "x"));
- val ts = map str
- (DatatypeProp.indexify_names (replicate (length dts) "t"));
- val (_, id) = get_const_id gr cname
- in
- mk_let
- (map2 (fn p => fn q => mk_tuple [p, q]) xs ts ~~ gs)
- (mk_tuple
- [case xs of
- _ :: _ :: _ => Pretty.block
- [str id, Pretty.brk 1, mk_tuple xs]
- | _ => mk_app false (str id) xs,
- mk_delay (Pretty.block (mk_constr_term cname Ts T
- (map (single o mk_force) ts)))])
- end;
-
- fun mk_choice [c] = mk_constr "(i-1)" false c
- | mk_choice cs = Pretty.block [str "one_of",
- Pretty.brk 1, Pretty.blk (1, str "[" ::
- List.concat (separate [str ",", Pretty.fbrk]
- (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
- [str "]"]), Pretty.brk 1, str "()"];
-
- val gs = maps (fn s =>
- let val s' = strip_tname s
- in [str (s' ^ "G"), str (s' ^ "T")] end) tvs;
- val gen_name = "gen_" ^ snd (get_type_id gr tname)
-
- in
- Pretty.blk (4, separate (Pretty.brk 1)
- (str (prfx ^ gen_name ^
- (if null cs1 then "" else "'")) :: gs @
- (if null cs1 then [] else [str "i"]) @
- [str "j"]) @
- [str " =", Pretty.brk 1] @
- (if not (null cs1) andalso not (null cs2)
- then [str "frequency", Pretty.brk 1,
- Pretty.blk (1, [str "[",
- mk_tuple [str "i", mk_delay (mk_choice cs1)],
- str ",", Pretty.fbrk,
- mk_tuple [str "1", mk_delay (mk_choice cs2)],
- str "]"]), Pretty.brk 1, str "()"]
- else if null cs2 then
- [Pretty.block [str "(case", Pretty.brk 1,
- str "i", Pretty.brk 1, str "of",
- Pretty.brk 1, str "0 =>", Pretty.brk 1,
- mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
- Pretty.brk 1, str "| _ =>", Pretty.brk 1,
- mk_choice cs1, str ")"]]
- else [mk_choice cs2])) ::
- (if null cs1 then []
- else [Pretty.blk (4, separate (Pretty.brk 1)
- (str ("and " ^ gen_name) :: gs @ [str "i"]) @
- [str " =", Pretty.brk 1] @
- separate (Pretty.brk 1) (str (gen_name ^ "'") :: gs @
- [str "i", str "i"]))]) @
- mk_gen_of_def gr "and " xs
- end
-
- in
- (module', (add_edge_acyclic (node_id, dep) gr
- handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
- let
- val gr1 = add_edge (node_id, dep)
- (new_node (node_id, (NONE, "", "")) gr);
- val (dtdef, gr2) = mk_dtdef "datatype " descr' gr1 ;
- in
- map_node node_id (K (NONE, module',
- string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
- [str ";"])) ^ "\n\n" ^
- (if "term_of" mem !mode then
- string_of (Pretty.blk (0, separate Pretty.fbrk
- (mk_term_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
- else "") ^
- (if "test" mem !mode then
- string_of (Pretty.blk (0, separate Pretty.fbrk
- (mk_gen_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
- else ""))) gr2
- end)
- end;
-
-
-(* case expressions *)
-
-fun pretty_case thy defs dep module brack constrs (c as Const (_, T)) ts gr =
- let val i = length constrs
- in if length ts <= i then
- invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr
- else
- let
- val ts1 = Library.take (i, ts);
- val t :: ts2 = Library.drop (i, ts);
- val names = List.foldr OldTerm.add_term_names
- (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1;
- val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
-
- fun pcase [] [] [] gr = ([], gr)
- | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr =
- let
- val j = length cargs;
- val xs = Name.variant_list names (replicate j "x");
- val Us' = Library.take (j, fst (strip_type U));
- val frees = map Free (xs ~~ Us');
- val (cp, gr0) = invoke_codegen thy defs dep module false
- (list_comb (Const (cname, Us' ---> dT), frees)) gr;
- val t' = Envir.beta_norm (list_comb (t, frees));
- val (p, gr1) = invoke_codegen thy defs dep module false t' gr0;
- val (ps, gr2) = pcase cs ts Us gr1;
- in
- ([Pretty.block [cp, str " =>", Pretty.brk 1, p]] :: ps, gr2)
- end;
-
- val (ps1, gr1) = pcase constrs ts1 Ts gr ;
- val ps = List.concat (separate [Pretty.brk 1, str "| "] ps1);
- val (p, gr2) = invoke_codegen thy defs dep module false t gr1;
- val (ps2, gr3) = fold_map (invoke_codegen thy defs dep module true) ts2 gr2;
- in ((if not (null ts2) andalso brack then parens else I)
- (Pretty.block (separate (Pretty.brk 1)
- (Pretty.block ([str "(case ", p, str " of",
- Pretty.brk 1] @ ps @ [str ")"]) :: ps2))), gr3)
- end
- end;
-
-
-(* constructors *)
-
-fun pretty_constr thy defs dep module brack args (c as Const (s, T)) ts gr =
- let val i = length args
- in if i > 1 andalso length ts < i then
- invoke_codegen thy defs dep module brack (eta_expand c ts i) gr
- else
- let
- val id = mk_qual_id module (get_const_id gr s);
- val (ps, gr') = fold_map
- (invoke_codegen thy defs dep module (i = 1)) ts gr;
- in (case args of
- _ :: _ :: _ => (if brack then parens else I)
- (Pretty.block [str id, Pretty.brk 1, mk_tuple ps])
- | _ => (mk_app brack (str id) ps), gr')
- end
- end;
-
-
-(* code generators for terms and types *)
-
-fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
- (c as Const (s, T), ts) =>
- (case DatatypePackage.datatype_of_case thy s of
- SOME {index, descr, ...} =>
- if is_some (get_assoc_code thy (s, T)) then NONE else
- SOME (pretty_case thy defs dep module brack
- (#3 (the (AList.lookup op = descr index))) c ts gr )
- | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
- (SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
- if is_some (get_assoc_code thy (s, T)) then NONE else
- let
- val SOME (tyname', _, constrs) = AList.lookup op = descr index;
- val SOME args = AList.lookup op = constrs s
- in
- if tyname <> tyname' then NONE
- else SOME (pretty_constr thy defs
- dep module brack args c ts (snd (invoke_tycodegen thy defs dep module false U gr)))
- end
- | _ => NONE)
- | _ => NONE);
-
-fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
- (case DatatypePackage.get_datatype thy s of
- NONE => NONE
- | SOME {descr, sorts, ...} =>
- if is_some (get_assoc_type thy s) then NONE else
- let
- val (ps, gr') = fold_map
- (invoke_tycodegen thy defs dep module false) Ts gr;
- val (module', gr'') = add_dt_defs thy defs dep module descr sorts gr' ;
- val (tyid, gr''') = mk_type_id module' s gr''
- in SOME (Pretty.block ((if null Ts then [] else
- [mk_tuple ps, str " "]) @
- [str (mk_qual_id module tyid)]), gr''')
- end)
- | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
-
-
-(** generic code generator **)
-
-(* liberal addition of code data for datatypes *)
-
-fun mk_constr_consts thy vs dtco cos =
- let
- val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
- val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
- in if is_some (try (Code.constrset_of_consts thy) cs')
- then SOME cs
- else NONE
- end;
-
-
-(* case certificates *)
-
-fun mk_case_cert thy tyco =
- let
- val raw_thms =
- (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
- val thms as hd_thm :: _ = raw_thms
- |> Conjunction.intr_balanced
- |> Thm.unvarify
- |> Conjunction.elim_balanced (length raw_thms)
- |> map Simpdata.mk_meta_eq
- |> map Drule.zero_var_indexes
- val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
- | _ => I) (Thm.prop_of hd_thm) [];
- val rhs = hd_thm
- |> Thm.prop_of
- |> Logic.dest_equals
- |> fst
- |> Term.strip_comb
- |> apsnd (fst o split_last)
- |> list_comb;
- val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
- val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
- in
- thms
- |> Conjunction.intr_balanced
- |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
- |> Thm.implies_intr asm
- |> Thm.generalize ([], params) 0
- |> AxClass.unoverload thy
- |> Thm.varifyT
- end;
-
-
-(* equality *)
-
-fun mk_eq_eqns thy dtco =
- let
- val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
- val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
- $ t1 $ t2;
- fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
- fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
- val triv_injects = map_filter
- (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
- | _ => NONE) cos;
- fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
- trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
- val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
- fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
- [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
- val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
- val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
- val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
- addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
- addsimprocs [DatatypePackage.distinct_simproc]);
- fun prove prop = SkipProof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
- |> Simpdata.mk_eq;
- in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
-
-fun add_equality vs dtcos thy =
- let
- fun add_def dtco lthy =
- let
- val ty = Type (dtco, map TFree vs);
- fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
- $ Free ("x", ty) $ Free ("y", ty);
- val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
- val def' = Syntax.check_term lthy def;
- val ((_, (_, thm)), lthy') = Specification.definition
- (NONE, (Attrib.empty_binding, def')) lthy;
- val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
- val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
- in (thm', lthy') end;
- fun tac thms = Class.intro_classes_tac []
- THEN ALLGOALS (ProofContext.fact_tac thms);
- fun add_eq_thms dtco thy =
- let
- val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
- val thy_ref = Theory.check_thy thy;
- fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
- in
- Code.add_eqnl (const, Lazy.lazy mk_thms) thy
- end;
- in
- thy
- |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
- |> fold_map add_def dtcos
- |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
- (fn _ => fn def_thms => tac def_thms) def_thms)
- |-> (fn def_thms => fold Code.del_eqn def_thms)
- |> fold add_eq_thms dtcos
- end;
-
-
-(* register a datatype etc. *)
-
-fun add_all_code config dtcos thy =
- let
- val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
- val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
- val css = if exists is_none any_css then []
- else map_filter I any_css;
- val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
- val certs = map (mk_case_cert thy) dtcos;
- in
- if null css then thy
- else thy
- |> tap (fn _ => DatatypeAux.message config "Registering datatype for code generator ...")
- |> fold Code.add_datatype css
- |> fold_rev Code.add_default_eqn case_rewrites
- |> fold Code.add_case certs
- |> add_equality vs dtcos
- end;
-
-
-(** theory setup **)
-
-val setup =
- add_codegen "datatype" datatype_codegen
- #> add_tycodegen "datatype" datatype_tycodegen
- #> DatatypePackage.interpretation add_all_code
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,705 +0,0 @@
-(* Title: HOL/Tools/datatype_package.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package for Isabelle/HOL.
-*)
-
-signature DATATYPE_PACKAGE =
-sig
- type datatype_config = DatatypeAux.datatype_config
- type datatype_info = DatatypeAux.datatype_info
- type descr = DatatypeAux.descr
- val get_datatypes : theory -> datatype_info Symtab.table
- val get_datatype : theory -> string -> datatype_info option
- val the_datatype : theory -> string -> datatype_info
- val datatype_of_constr : theory -> string -> datatype_info option
- val datatype_of_case : theory -> string -> datatype_info option
- val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
- val the_datatype_descr : theory -> string list
- -> descr * (string * sort) list * string list
- * (string list * string list) * (typ list * typ list)
- val get_datatype_constrs : theory -> string -> (string * typ) list option
- val distinct_simproc : simproc
- val make_case : Proof.context -> bool -> string list -> term ->
- (term * term) list -> term * (term * (int * bool)) list
- val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
- val read_typ: theory ->
- (typ list * (string * sort) list) * string -> typ list * (string * sort) list
- val interpretation : (datatype_config -> string list -> theory -> theory) -> theory -> theory
- type rules = {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list}
- val rep_datatype : datatype_config -> (rules -> Proof.context -> Proof.context)
- -> string list option -> term list -> theory -> Proof.state;
- val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
- val add_datatype : datatype_config -> string list -> (string list * binding * mixfix *
- (binding * typ list * mixfix) list) list -> theory -> rules * theory
- val add_datatype_cmd : string list -> (string list * binding * mixfix *
- (binding * string list * mixfix) list) list -> theory -> rules * theory
- val setup: theory -> theory
- val print_datatypes : theory -> unit
-end;
-
-structure DatatypePackage : DATATYPE_PACKAGE =
-struct
-
-open DatatypeAux;
-
-
-(* theory data *)
-
-structure DatatypesData = TheoryDataFun
-(
- type T =
- {types: datatype_info Symtab.table,
- constrs: datatype_info Symtab.table,
- cases: datatype_info Symtab.table};
-
- val empty =
- {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
- val copy = I;
- val extend = I;
- fun merge _
- ({types = types1, constrs = constrs1, cases = cases1},
- {types = types2, constrs = constrs2, cases = cases2}) =
- {types = Symtab.merge (K true) (types1, types2),
- constrs = Symtab.merge (K true) (constrs1, constrs2),
- cases = Symtab.merge (K true) (cases1, cases2)};
-);
-
-val get_datatypes = #types o DatatypesData.get;
-val map_datatypes = DatatypesData.map;
-
-fun print_datatypes thy =
- Pretty.writeln (Pretty.strs ("datatypes:" ::
- map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
-
-
-(** theory information about datatypes **)
-
-fun put_dt_infos (dt_infos : (string * datatype_info) list) =
- map_datatypes (fn {types, constrs, cases} =>
- {types = fold Symtab.update dt_infos types,
- constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
- (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
- (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
- cases = fold Symtab.update
- (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
- cases});
-
-val get_datatype = Symtab.lookup o get_datatypes;
-
-fun the_datatype thy name = (case get_datatype thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
-
-val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
-val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
-
-fun get_datatype_descr thy dtco =
- get_datatype thy dtco
- |> Option.map (fn info as { descr, index, ... } =>
- (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
-
-fun the_datatype_spec thy dtco =
- let
- val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
- val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
- o DatatypeAux.dest_DtTFree) dtys;
- val cos = map
- (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
- in (sorts, cos) end;
-
-fun the_datatype_descr thy (raw_tycos as raw_tyco :: _) =
- let
- val info = the_datatype thy raw_tyco;
- val descr = #descr info;
-
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
- val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
- o dest_DtTFree) dtys;
-
- fun is_DtTFree (DtTFree _) = true
- | is_DtTFree _ = false
- val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
- val protoTs as (dataTs, _) = chop k descr
- |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
-
- val tycos = map fst dataTs;
- val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
- else error ("Type constructors " ^ commas (map quote raw_tycos)
- ^ "do not belong exhaustively to one mutual recursive datatype");
-
- val (Ts, Us) = (pairself o map) Type protoTs;
-
- val names = map Long_Name.base_name (the_default tycos (#alt_names info));
- val (auxnames, _) = Name.make_context names
- |> fold_map (yield_singleton Name.variants o name_of_typ) Us
-
- in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
-
-fun get_datatype_constrs thy dtco =
- case try (the_datatype_spec thy) dtco
- of SOME (sorts, cos) =>
- let
- fun subst (v, sort) = TVar ((v, 0), sort);
- fun subst_ty (TFree v) = subst v
- | subst_ty ty = ty;
- val dty = Type (dtco, map subst sorts);
- fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
- in SOME (map mk_co cos) end
- | NONE => NONE;
-
-
-(** induct method setup **)
-
-(* case names *)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = maps dt_recs dts
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct (op =) (maps dt_recs args));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
-
-fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (RuleCases.case_names o exhaust_cases descr o #1)
- (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
-
-end;
-
-fun add_rules simps case_thms rec_thms inject distinct
- weak_case_congs cong_att =
- PureThy.add_thmss [((Binding.name "simps", simps), []),
- ((Binding.empty, flat case_thms @
- flat distinct @ rec_thms), [Simplifier.simp_add]),
- ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
- ((Binding.empty, flat inject), [iff_add]),
- ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
- ((Binding.empty, weak_case_congs), [cong_att])]
- #> snd;
-
-
-(* add_cases_induct *)
-
-fun add_cases_induct infos induction thy =
- let
- val inducts = ProjectRule.projections (ProofContext.init thy) induction;
-
- fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
- [((Binding.empty, nth inducts index), [Induct.induct_type name]),
- ((Binding.empty, exhaustion), [Induct.cases_type name])];
- fun unnamed_rule i =
- ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
- in
- thy |> PureThy.add_thms
- (maps named_rules infos @
- map unnamed_rule (length infos upto length inducts - 1)) |> snd
- |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
- end;
-
-
-
-(**** simplification procedure for showing distinctness of constructors ****)
-
-fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
- | stripT p = p;
-
-fun stripC (i, f $ x) = stripC (i + 1, f)
- | stripC p = p;
-
-val distinctN = "constr_distinct";
-
-fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
- FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
- atac 2, resolve_tac thms 1, etac FalseE 1]))
- | ManyConstrs (thm, simpset) =>
- let
- val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
- map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
- ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
- in
- Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
- full_simp_tac (Simplifier.inherit_context ss simpset) 1,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
- etac FalseE 1]))
- end;
-
-fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
- (case (stripC (0, t1), stripC (0, t2)) of
- ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
- (case (stripT (0, T1), stripT (0, T2)) of
- ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
- if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
- (case (get_datatype_descr thy) tname1 of
- SOME (_, (_, constrs)) => let val cnames = map fst constrs
- in if cname1 mem cnames andalso cname2 mem cnames then
- SOME (distinct_rule thy ss tname1
- (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
- else NONE
- end
- | NONE => NONE)
- else NONE
- | _ => NONE)
- | _ => NONE)
- | distinct_proc _ _ _ = NONE;
-
-val distinct_simproc =
- Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
-
-val dist_ss = HOL_ss addsimprocs [distinct_simproc];
-
-val simproc_setup =
- Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
-
-
-(**** translation rules for case ****)
-
-fun make_case ctxt = DatatypeCase.make_case
- (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
-
-fun strip_case ctxt = DatatypeCase.strip_case
- (datatype_of_case (ProofContext.theory_of ctxt));
-
-fun add_case_tr' case_names thy =
- Sign.add_advanced_trfuns ([], [],
- map (fn case_name =>
- let val case_name' = Sign.const_syntax_name thy case_name
- in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
- end) case_names, []) thy;
-
-val trfun_setup =
- Sign.add_advanced_trfuns ([],
- [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
- [], []);
-
-
-(* prepare types *)
-
-fun read_typ thy ((Ts, sorts), str) =
- let
- val ctxt = ProofContext.init thy
- |> fold (Variable.declare_typ o TFree) sorts;
- val T = Syntax.read_typ ctxt str;
- in (Ts @ [T], Term.add_tfreesT T sorts) end;
-
-fun cert_typ sign ((Ts, sorts), raw_T) =
- let
- val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
- TYPE (msg, _, _) => error msg;
- val sorts' = Term.add_tfreesT T sorts;
- in (Ts @ [T],
- case duplicates (op =) (map fst sorts') of
- [] => sorts'
- | dups => error ("Inconsistent sort constraints for " ^ commas dups))
- end;
-
-
-(**** make datatype info ****)
-
-fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
- (((((((((i, (_, (tname, _, _))), case_name), case_thms),
- exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
- (tname,
- {index = i,
- alt_names = alt_names,
- descr = descr,
- sorts = sorts,
- rec_names = reccomb_names,
- rec_rewrites = rec_thms,
- case_name = case_name,
- case_rewrites = case_thms,
- induction = induct,
- exhaustion = exhaustion_thm,
- distinct = distinct_thm,
- inject = inject,
- nchotomy = nchotomy,
- case_cong = case_cong,
- weak_case_cong = weak_case_cong});
-
-type rules = {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list}
-
-structure DatatypeInterpretation = InterpretationFun
- (type T = datatype_config * string list val eq: T * T -> bool = eq_snd op =);
-fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
-
-
-(******************* definitional introduction of datatypes *******************)
-
-fun add_datatype_def (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy =
- let
- val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
- DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts
- types_syntax constr_syntax case_names_induct;
-
- val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names descr
- sorts induct case_names_exhausts thy2;
- val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
- config new_type_names descr sorts dt_info inject dist_rewrites
- (Simplifier.theory_context thy3 dist_ss) induct thy3;
- val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
- config new_type_names descr sorts reccomb_names rec_thms thy4;
- val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms config new_type_names
- descr sorts inject dist_rewrites casedist_thms case_thms thy6;
- val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys config new_type_names
- descr sorts casedist_thms thy7;
- val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
- descr sorts nchotomys case_thms thy8;
- val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- descr sorts thy9;
-
- val dt_infos = map
- (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
- ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
- casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy12 =
- thy10
- |> add_case_tr' case_names
- |> Sign.add_path (space_implode "_" new_type_names)
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
- |> DatatypeInterpretation.data (config, map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct,
- simps = simps}, thy12)
- end;
-
-
-(*********************** declare existing type as datatype *********************)
-
-fun prove_rep_datatype (config : datatype_config) alt_names new_type_names descr sorts induct inject half_distinct thy =
- let
- val ((_, [induct']), _) =
- Variable.importT_thms [induct] (Variable.thm_context induct);
-
- fun err t = error ("Ill-formed predicate in induction rule: " ^
- Syntax.string_of_term_global thy t);
-
- fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
- ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
- | get_typ t = err t;
- val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
-
- val dt_info = get_datatypes thy;
-
- val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
- val (case_names_induct, case_names_exhausts) =
- (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
-
- val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val (casedist_thms, thy2) = thy |>
- DatatypeAbsProofs.prove_casedist_thms config new_type_names [descr] sorts induct
- case_names_exhausts;
- val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
- config new_type_names [descr] sorts dt_info inject distinct
- (Simplifier.theory_context thy2 dist_ss) induct thy2;
- val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms
- config new_type_names [descr] sorts reccomb_names rec_thms thy3;
- val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
- config new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
- val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys config new_type_names
- [descr] sorts casedist_thms thy5;
- val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
- [descr] sorts nchotomys case_thms thy6;
- val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- [descr] sorts thy7;
-
- val ((_, [induct']), thy10) =
- thy8
- |> store_thmss "inject" new_type_names inject
- ||>> store_thmss "distinct" new_type_names distinct
- ||> Sign.add_path (space_implode "_" new_type_names)
- ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
-
- val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
- ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
- map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy11 =
- thy10
- |> add_case_tr' case_names
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct'
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
- |> snd
- |> DatatypeInterpretation.data (config, map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct',
- simps = simps}, thy11)
- end;
-
-fun gen_rep_datatype prep_term (config : datatype_config) after_qed alt_names raw_ts thy =
- let
- fun constr_of_term (Const (c, T)) = (c, T)
- | constr_of_term t =
- error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
- fun no_constr (c, T) = error ("Bad constructor: "
- ^ Sign.extern_const thy c ^ "::"
- ^ Syntax.string_of_typ_global thy T);
- fun type_of_constr (cT as (_, T)) =
- let
- val frees = OldTerm.typ_tfrees T;
- val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
- handle TYPE _ => no_constr cT
- val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
- val _ = if length frees <> length vs then no_constr cT else ();
- in (tyco, (vs, cT)) end;
-
- val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
- val _ = case map_filter (fn (tyco, _) =>
- if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
- of [] => ()
- | tycos => error ("Type(s) " ^ commas (map quote tycos)
- ^ " already represented inductivly");
- val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
- val ms = case distinct (op =) (map length raw_vss)
- of [n] => 0 upto n - 1
- | _ => error ("Different types in given constructors");
- fun inter_sort m = map (fn xs => nth xs m) raw_vss
- |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
- val sorts = map inter_sort ms;
- val vs = Name.names Name.context Name.aT sorts;
-
- fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
- (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
-
- val cs = map (apsnd (map norm_constr)) raw_cs;
- val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
- o fst o strip_type;
- val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
-
- fun mk_spec (i, (tyco, constr)) = (i, (tyco,
- map (DtTFree o fst) vs,
- (map o apsnd) dtyps_of_typ constr))
- val descr = map_index mk_spec cs;
- val injs = DatatypeProp.make_injs [descr] vs;
- val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
- val ind = DatatypeProp.make_ind [descr] vs;
- val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
-
- fun after_qed' raw_thms =
- let
- val [[[induct]], injs, half_distincts] =
- unflat rules (map Drule.zero_var_indexes_list raw_thms);
- (*FIXME somehow dubious*)
- in
- ProofContext.theory_result
- (prove_rep_datatype config alt_names new_type_names descr vs induct injs half_distincts)
- #-> after_qed
- end;
- in
- thy
- |> ProofContext.init
- |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
- end;
-
-val rep_datatype = gen_rep_datatype Sign.cert_term;
-val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_datatype_config (K I);
-
-
-
-(******************************** add datatype ********************************)
-
-fun gen_add_datatype prep_typ (config : datatype_config) new_type_names dts thy =
- let
- val _ = Theory.requires thy "Datatype" "datatype definitions";
-
- (* this theory is used just for parsing *)
-
- val tmp_thy = thy |>
- Theory.copy |>
- Sign.add_types (map (fn (tvs, tname, mx, _) =>
- (tname, length tvs, mx)) dts);
-
- val (tyvars, _, _, _)::_ = dts;
- val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
- let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
- in (case duplicates (op =) tvs of
- [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
- else error ("Mutually recursive datatypes must have same type parameters")
- | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
- " : " ^ commas dups))
- end) dts);
-
- val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
- [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
-
- fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
- let
- fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
- let
- val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
- val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
- [] => ()
- | vs => error ("Extra type variables on rhs: " ^ commas vs))
- in (constrs @ [((if #flat_names config then Sign.full_name tmp_thy else
- Sign.full_name_path tmp_thy tname')
- (Binding.map_name (Syntax.const_name mx') cname),
- map (dtyp_of_typ new_dts) cargs')],
- constr_syntax' @ [(cname, mx')], sorts'')
- end handle ERROR msg => cat_error msg
- ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
- " of datatype " ^ quote (Binding.str_of tname));
-
- val (constrs', constr_syntax', sorts') =
- fold prep_constr constrs ([], [], sorts)
-
- in
- case duplicates (op =) (map fst constrs') of
- [] =>
- (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
- map DtTFree tvs, constrs'))],
- constr_syntax @ [constr_syntax'], sorts', i + 1)
- | dups => error ("Duplicate constructors " ^ commas dups ^
- " in datatype " ^ quote (Binding.str_of tname))
- end;
-
- val (dts', constr_syntax, sorts', i) =
- fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
- val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
- val dt_info = get_datatypes thy;
- val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
- val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
- if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
- else raise exn;
-
- val descr' = flat descr;
- val case_names_induct = mk_case_names_induct descr';
- val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
- in
- add_datatype_def
- (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy
- end;
-
-val add_datatype = gen_add_datatype cert_typ;
-val add_datatype_cmd = gen_add_datatype read_typ default_datatype_config;
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- DatatypeRepProofs.distinctness_limit_setup #>
- simproc_setup #>
- trfun_setup #>
- DatatypeInterpretation.init;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
- (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
-
-fun mk_datatype args =
- let
- val names = map
- (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
- val specs = map (fn ((((_, vs), t), mx), cons) =>
- (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
- in snd o add_datatype_cmd names specs end;
-
-val _ =
- OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-val _ =
- OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
- (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
- >> (fn (alt_names, ts) => Toplevel.print
- o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
-
-end;
-
-
-(* document antiquotation *)
-
-val _ = ThyOutput.antiquotation "datatype" Args.tyname
- (fn {source = src, context = ctxt, ...} => fn dtco =>
- let
- val thy = ProofContext.theory_of ctxt;
- val (vs, cos) = the_datatype_spec thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
- Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
- | pretty_typ_bracket ty =
- Syntax.pretty_typ ctxt ty;
- fun pretty_constr (co, tys) =
- (Pretty.block o Pretty.breaks)
- (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
- map pretty_typ_bracket tys);
- val pretty_datatype =
- Pretty.block
- (Pretty.command "datatype" :: Pretty.brk 1 ::
- Syntax.pretty_typ ctxt ty ::
- Pretty.str " =" :: Pretty.brk 1 ::
- flat (separate [Pretty.brk 1, Pretty.str "| "]
- (map (single o pretty_constr) cos)));
- in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
-
-end;
-
--- a/src/HOL/Tools/datatype_package/datatype_prop.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,435 +0,0 @@
-(* Title: HOL/Tools/datatype_prop.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Characteristic properties of datatypes.
-*)
-
-signature DATATYPE_PROP =
-sig
- val indexify_names: string list -> string list
- val make_tnames: typ list -> string list
- val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list
- val make_distincts : DatatypeAux.descr list ->
- (string * sort) list -> (int * term list) list (*no symmetric inequalities*)
- val make_ind : DatatypeAux.descr list -> (string * sort) list -> term
- val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list
- val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list ->
- string list -> typ list * typ list
- val make_primrecs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_cases : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list list
- val make_splits : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> (term * term) list
- val make_weak_case_congs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_case_congs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_nchotomys : DatatypeAux.descr list ->
- (string * sort) list -> term list
-end;
-
-structure DatatypeProp : DATATYPE_PROP =
-struct
-
-open DatatypeAux;
-
-fun indexify_names names =
- let
- fun index (x :: xs) tab =
- (case AList.lookup (op =) tab x of
- NONE => if member (op =) xs x then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
- | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
- | index [] _ = [];
- in index names [] end;
-
-fun make_tnames Ts =
- let
- fun type_name (TFree (name, _)) = implode (tl (explode name))
- | type_name (Type (name, _)) =
- let val name' = Long_Name.base_name name
- in if Syntax.is_identifier name' then name' else "x" end;
- in indexify_names (map type_name Ts) end;
-
-
-(************************* injectivity of constructors ************************)
-
-fun make_injs descr sorts =
- let
- val descr' = flat descr;
- fun make_inj T (cname, cargs) =
- if null cargs then I else
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val constr_t = Const (cname, Ts ---> T);
- val tnames = make_tnames Ts;
- val frees = map Free (tnames ~~ Ts);
- val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
- in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
- foldr1 (HOLogic.mk_binop "op &")
- (map HOLogic.mk_eq (frees ~~ frees')))))
- end;
- in
- map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
- (hd descr) (Library.take (length (hd descr), get_rec_types descr' sorts))
- end;
-
-
-(************************* distinctness of constructors ***********************)
-
-fun make_distincts descr sorts =
- let
- val descr' = flat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun prep_constr (cname, cargs) = (cname, map (typ_of_dtyp descr' sorts) cargs);
-
- fun make_distincts' _ [] = []
- | make_distincts' T ((cname, cargs)::constrs) =
- let
- val frees = map Free ((make_tnames cargs) ~~ cargs);
- val t = list_comb (Const (cname, cargs ---> T), frees);
-
- fun make_distincts'' (cname', cargs') =
- let
- val frees' = map Free ((map ((op ^) o (rpair "'"))
- (make_tnames cargs')) ~~ cargs');
- val t' = list_comb (Const (cname', cargs' ---> T), frees')
- in
- HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.mk_eq (t, t'))
- end
-
- in map make_distincts'' constrs @ make_distincts' T constrs end;
-
- in
- map2 (fn ((_, (_, _, constrs))) => fn T =>
- (length constrs, make_distincts' T (map prep_constr constrs))) (hd descr) newTs
- end;
-
-
-(********************************* induction **********************************)
-
-fun make_ind descr sorts =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val pnames = if length descr' = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr');
-
- fun make_pred i T =
- let val T' = T --> HOLogic.boolT
- in Free (List.nth (pnames, i), T') end;
-
- fun make_ind_prem k T (cname, cargs) =
- let
- fun mk_prem ((dt, s), T) =
- let val (Us, U) = strip_type T
- in list_all (map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us)))
- end;
-
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr' sorts) recs;
- val tnames = Name.variant_list pnames (make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = tnames ~~ Ts;
- val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
-
- in list_all_free (frees, Logic.list_implies (prems,
- HOLogic.mk_Trueprop (make_pred k T $
- list_comb (Const (cname, Ts ---> T), map Free frees))))
- end;
-
- val prems = List.concat (map (fn ((i, (_, _, constrs)), T) =>
- map (make_ind_prem i T) constrs) (descr' ~~ recTs));
- val tnames = make_tnames recTs;
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
- (descr' ~~ recTs ~~ tnames)))
-
- in Logic.list_implies (prems, concl) end;
-
-(******************************* case distinction *****************************)
-
-fun make_casedists descr sorts =
- let
- val descr' = List.concat descr;
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees
- in list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
- HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))))
- end;
-
- fun make_casedist ((_, (_, _, constrs)), T) =
- let val prems = map (make_casedist_prem T) constrs
- in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)))
- end
-
- in map make_casedist
- ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts))
- end;
-
-(*************** characteristic equations for primrec combinator **************)
-
-fun make_primrec_Ts descr sorts used =
- let
- val descr' = List.concat descr;
-
- val rec_result_Ts = map TFree (Name.variant_list used (replicate (length descr') "'t") ~~
- replicate (length descr') HOLogic.typeS);
-
- val reccomb_fn_Ts = List.concat (map (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts);
-
- fun mk_argT (dt, T) =
- binder_types T ---> List.nth (rec_result_Ts, body_index dt);
-
- val argTs = Ts @ map mk_argT recs
- in argTs ---> List.nth (rec_result_Ts, i)
- end) constrs) descr');
-
- in (rec_result_Ts, reccomb_fn_Ts) end;
-
-fun make_primrecs new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
-
- val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.intern_const thy)
- (if length descr' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) =
- let
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr' sorts) recs;
- val tnames = make_tnames Ts;
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = map Free (tnames ~~ Ts);
- val frees' = map Free (rec_tnames ~~ recTs');
-
- fun mk_reccomb ((dt, T), t) =
- let val (Us, U) = strip_type T
- in list_abs (map (pair "x") Us,
- List.nth (reccombs, body_index dt) $ app_bnds t (length Us))
- end;
-
- val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees')
-
- in (ts @ [HOLogic.mk_Trueprop (HOLogic.mk_eq
- (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
- list_comb (f, frees @ reccombs')))], fs)
- end
-
- in fst (Library.foldl (fn (x, ((dt, T), comb_t)) =>
- Library.foldl (make_primrec T comb_t) (x, #3 (snd dt)))
- (([], rec_fns), descr' ~~ recTs ~~ reccombs))
- end;
-
-(****************** make terms of form t_case f1 ... fn *********************)
-
-fun make_case_combs new_type_names descr sorts thy fname =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used "'t", HOLogic.typeS);
-
- val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let val Ts = map (typ_of_dtyp descr' sorts) cargs
- in Ts ---> T' end) constrs) (hd descr);
-
- val case_names = map (fn s =>
- Sign.intern_const thy (s ^ "_case")) new_type_names
- in
- map (fn ((name, Ts), T) => list_comb
- (Const (name, Ts @ [T] ---> T'),
- map (uncurry (mk_Free fname)) (Ts ~~ (1 upto length Ts))))
- (case_names ~~ case_fn_Ts ~~ newTs)
- end;
-
-(**************** characteristic equations for case combinator ****************)
-
-fun make_cases new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun make_case T comb_t ((cname, cargs), f) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = map Free ((make_tnames Ts) ~~ Ts)
- in HOLogic.mk_Trueprop (HOLogic.mk_eq
- (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
- list_comb (f, frees)))
- end
-
- in map (fn (((_, (_, _, constrs)), T), comb_t) =>
- map (make_case T comb_t) (constrs ~~ (snd (strip_comb comb_t))))
- ((hd descr) ~~ newTs ~~ (make_case_combs new_type_names descr sorts thy "f"))
- end;
-
-
-(*************************** the "split" - equations **************************)
-
-fun make_splits new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used' = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
- val P = Free ("P", T' --> HOLogic.boolT);
-
- fun make_split (((_, (_, _, constrs)), T), comb_t) =
- let
- val (_, fs) = strip_comb comb_t;
- val used = ["P", "x"] @ (map (fst o dest_Free) fs);
-
- fun process_constr (((cname, cargs), f), (t1s, t2s)) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts);
- val eqn = HOLogic.mk_eq (Free ("x", T),
- list_comb (Const (cname, Ts ---> T), frees));
- val P' = P $ list_comb (f, frees)
- in ((List.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t))
- (HOLogic.imp $ eqn $ P') frees)::t1s,
- (List.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t))
- (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) frees)::t2s)
- end;
-
- val (t1s, t2s) = List.foldr process_constr ([], []) (constrs ~~ fs);
- val lhs = P $ (comb_t $ Free ("x", T))
- in
- (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ mk_disj t2s)))
- end
-
- in map make_split ((hd descr) ~~ newTs ~~
- (make_case_combs new_type_names descr sorts thy "f"))
- end;
-
-(************************* additional rules for TFL ***************************)
-
-fun make_weak_case_congs new_type_names descr sorts thy =
- let
- val case_combs = make_case_combs new_type_names descr sorts thy "f";
-
- fun mk_case_cong comb =
- let
- val Type ("fun", [T, _]) = fastype_of comb;
- val M = Free ("M", T);
- val M' = Free ("M'", T);
- in
- Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb $ M')))
- end
- in
- map mk_case_cong case_combs
- end;
-
-
-(*---------------------------------------------------------------------------
- * Structure of case congruence theorem looks like this:
- *
- * (M = M')
- * ==> (!!x1,...,xk. (M' = C1 x1..xk) ==> (f1 x1..xk = g1 x1..xk))
- * ==> ...
- * ==> (!!x1,...,xj. (M' = Cn x1..xj) ==> (fn x1..xj = gn x1..xj))
- * ==>
- * (ty_case f1..fn M = ty_case g1..gn M')
- *---------------------------------------------------------------------------*)
-
-fun make_case_congs new_type_names descr sorts thy =
- let
- val case_combs = make_case_combs new_type_names descr sorts thy "f";
- val case_combs' = make_case_combs new_type_names descr sorts thy "g";
-
- fun mk_case_cong ((comb, comb'), (_, (_, _, constrs))) =
- let
- val Type ("fun", [T, _]) = fastype_of comb;
- val (_, fs) = strip_comb comb;
- val (_, gs) = strip_comb comb';
- val used = ["M", "M'"] @ map (fst o dest_Free) (fs @ gs);
- val M = Free ("M", T);
- val M' = Free ("M'", T);
-
- fun mk_clause ((f, g), (cname, _)) =
- let
- val (Ts, _) = strip_type (fastype_of f);
- val tnames = Name.variant_list used (make_tnames Ts);
- val frees = map Free (tnames ~~ Ts)
- in
- list_all_free (tnames ~~ Ts, Logic.mk_implies
- (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (M', list_comb (Const (cname, Ts ---> T), frees))),
- HOLogic.mk_Trueprop
- (HOLogic.mk_eq (list_comb (f, frees), list_comb (g, frees)))))
- end
-
- in
- Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')) ::
- map mk_clause (fs ~~ gs ~~ constrs),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb' $ M')))
- end
-
- in
- map mk_case_cong (case_combs ~~ case_combs' ~~ hd descr)
- end;
-
-(*---------------------------------------------------------------------------
- * Structure of exhaustion theorem looks like this:
- *
- * !v. (? y1..yi. v = C1 y1..yi) | ... | (? y1..yj. v = Cn y1..yj)
- *---------------------------------------------------------------------------*)
-
-fun make_nchotomys descr sorts =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun mk_eqn T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val tnames = Name.variant_list ["v"] (make_tnames Ts);
- val frees = tnames ~~ Ts
- in
- List.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t))
- (HOLogic.mk_eq (Free ("v", T),
- list_comb (Const (cname, Ts ---> T), map Free frees))) frees
- end
-
- in map (fn ((_, (_, _, constrs)), T) =>
- HOLogic.mk_Trueprop (HOLogic.mk_all ("v", T, mk_disj (map (mk_eqn T) constrs))))
- (hd descr ~~ newTs)
- end;
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_realizer.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,230 +0,0 @@
-(* Title: HOL/Tools/datatype_realizer.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Porgram extraction from proofs involving datatypes:
-Realizers for induction and case analysis
-*)
-
-signature DATATYPE_REALIZER =
-sig
- val add_dt_realizers: DatatypeAux.datatype_config -> string list -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure DatatypeRealizer : DATATYPE_REALIZER =
-struct
-
-open DatatypeAux;
-
-fun subsets i j = if i <= j then
- let val is = subsets (i+1) j
- in map (fn ks => i::ks) is @ is end
- else [[]];
-
-fun forall_intr_prf (t, prf) =
- let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
-
-fun prf_of thm =
- Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
-
-fun prf_subst_vars inst =
- Proofterm.map_proof_terms (subst_vars ([], inst)) I;
-
-fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
-
-fun tname_of (Type (s, _)) = s
- | tname_of _ = "";
-
-fun mk_realizes T = Const ("realizes", T --> HOLogic.boolT --> HOLogic.boolT);
-
-fun make_ind sorts ({descr, rec_names, rec_rewrites, induction, ...} : datatype_info) is thy =
- let
- val recTs = get_rec_types descr sorts;
- val pnames = if length descr = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
-
- val rec_result_Ts = map (fn ((i, _), P) =>
- if i mem is then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
- (descr ~~ pnames);
-
- fun make_pred i T U r x =
- if i mem is then
- Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
- else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
-
- fun mk_all i s T t =
- if i mem is then list_all_free ([(s, T)], t) else t;
-
- val (prems, rec_fns) = split_list (flat (fst (fold_map
- (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
- val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
- val frees = tnames ~~ Ts;
-
- fun mk_prems vs [] =
- let
- val rT = nth (rec_result_Ts) i;
- val vs' = filter_out is_unit vs;
- val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
- val f' = Envir.eta_contract (list_abs_free
- (map dest_Free vs, if i mem is then list_comb (f, vs')
- else HOLogic.unit));
- in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
- (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
- end
- | mk_prems vs (((dt, s), T) :: ds) =
- let
- val k = body_index dt;
- val (Us, U) = strip_type T;
- val i = length Us;
- val rT = nth (rec_result_Ts) k;
- val r = Free ("r" ^ s, Us ---> rT);
- val (p, f) = mk_prems (vs @ [r]) ds
- in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
- (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred k rT U (app_bnds r i)
- (app_bnds (Free (s, T)) i))), p)), f)
- end
-
- in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
- constrs) (descr ~~ recTs) 1)));
-
- fun mk_proj j [] t = t
- | mk_proj j (i :: is) t = if null is then t else
- if (j: int) = i then HOLogic.mk_fst t
- else mk_proj j is (HOLogic.mk_snd t);
-
- val tnames = DatatypeProp.make_tnames recTs;
- val fTs = map fastype_of rec_fns;
- val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
- (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
- (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
- val r = if null is then Extraction.nullt else
- foldr1 HOLogic.mk_prod (List.mapPartial (fn (((((i, _), T), U), s), tname) =>
- if i mem is then SOME
- (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
- else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), U), tname) =>
- make_pred i U T (mk_proj i is r) (Free (tname, T)))
- (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
- val cert = cterm_of thy;
- val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induction))) ~~ ps);
-
- val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
- (fn prems =>
- [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
- rtac (cterm_instantiate inst induction) 1,
- ALLGOALS ObjectLogic.atomize_prems_tac,
- rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
- REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
- REPEAT (etac allE i) THEN atac i)) 1)]);
-
- val ind_name = Thm.get_name induction;
- val vs = map (fn i => List.nth (pnames, i)) is;
- val (thm', thy') = thy
- |> Sign.root_path
- |> PureThy.store_thm
- (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
- ||> Sign.restore_naming thy;
-
- val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []);
- val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
- val ivs1 = map Var (filter_out (fn (_, T) =>
- tname_of (body_type T) mem ["set", "bool"]) ivs);
- val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (AList.lookup (op =) rvs ixn))) ivs;
-
- val prf = List.foldr forall_intr_prf
- (List.foldr (fn ((f, p), prf) =>
- (case head_of (strip_abs_body f) of
- Free (s, T) =>
- let val T' = Logic.varifyT T
- in Abst (s, SOME T', Proofterm.prf_abstract_over
- (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
- end
- | _ => AbsP ("H", SOME p, prf)))
- (Proofterm.proof_combP
- (prf_of thm', map PBound (length prems - 1 downto 0))) (rec_fns ~~ prems_of thm)) ivs2;
-
- val r' = if null is then r else Logic.varify (List.foldr (uncurry lambda)
- r (map Logic.unvarify ivs1 @ filter_out is_unit
- (map (head_of o strip_abs_body) rec_fns)));
-
- in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
-
-
-fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaustion, ...} : datatype_info) thy =
- let
- val cert = cterm_of thy;
- val rT = TFree ("'P", HOLogic.typeS);
- val rT' = TVar (("'P", 0), HOLogic.typeS);
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees;
- val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
- in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
- HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
- list_comb (r, free_ts)))))
- end;
-
- val SOME (_, _, constrs) = AList.lookup (op =) descr index;
- val T = List.nth (get_rec_types descr sorts, index);
- val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
- val r = Const (case_name, map fastype_of rs ---> T --> rT);
-
- val y = Var (("y", 0), Logic.legacy_varifyT T);
- val y' = Free ("y", T);
-
- val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
- HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
- list_comb (r, rs @ [y'])))))
- (fn prems =>
- [rtac (cterm_instantiate [(cert y, cert y')] exhaustion) 1,
- ALLGOALS (EVERY'
- [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
- resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
-
- val exh_name = Thm.get_name exhaustion;
- val (thm', thy') = thy
- |> Sign.root_path
- |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
- ||> Sign.restore_naming thy;
-
- val P = Var (("P", 0), rT' --> HOLogic.boolT);
- val prf = forall_intr_prf (y, forall_intr_prf (P,
- List.foldr (fn ((p, r), prf) =>
- forall_intr_prf (Logic.legacy_varify r, AbsP ("H", SOME (Logic.varify p),
- prf))) (Proofterm.proof_combP (prf_of thm',
- map PBound (length prems - 1 downto 0))) (prems ~~ rs)));
- val r' = Logic.legacy_varify (Abs ("y", Logic.legacy_varifyT T,
- list_abs (map dest_Free rs, list_comb (r,
- map Bound ((length rs - 1 downto 0) @ [length rs])))));
-
- in Extraction.add_realizers_i
- [(exh_name, (["P"], r', prf)),
- (exh_name, ([], Extraction.nullt, prf_of exhaustion))] thy'
- end;
-
-fun add_dt_realizers config names thy =
- if ! Proofterm.proofs < 2 then thy
- else let
- val _ = message config "Adding realizers for induction and case analysis ..."
- val infos = map (DatatypePackage.the_datatype thy) names;
- val info :: _ = infos;
- in
- thy
- |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
- |> fold_rev (make_casedists (#sorts info)) infos
- end;
-
-val setup = DatatypePackage.interpretation add_dt_realizers;
-
-end;
--- a/src/HOL/Tools/datatype_package/datatype_rep_proofs.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,645 +0,0 @@
-(* Title: HOL/Tools/datatype_rep_proofs.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Definitional introduction of datatypes
-Proof of characteristic theorems:
-
- - injectivity of constructors
- - distinctness of constructors
- - induction theorem
-*)
-
-signature DATATYPE_REP_PROOFS =
-sig
- type datatype_config = DatatypeAux.datatype_config
- type descr = DatatypeAux.descr
- type datatype_info = DatatypeAux.datatype_info
- val distinctness_limit : int Config.T
- val distinctness_limit_setup : theory -> theory
- val representation_proofs : datatype_config -> datatype_info Symtab.table ->
- string list -> descr list -> (string * sort) list ->
- (binding * mixfix) list -> (binding * mixfix) list list -> attribute
- -> theory -> (thm list list * thm list list * thm list list *
- DatatypeAux.simproc_dist list * thm) * theory
-end;
-
-structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
-struct
-
-open DatatypeAux;
-
-(*the kind of distinctiveness axioms depends on number of constructors*)
-val (distinctness_limit, distinctness_limit_setup) =
- Attrib.config_int "datatype_distinctness_limit" 7;
-
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
-
-val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
-
-
-(** theory context references **)
-
-val f_myinv_f = thm "f_myinv_f";
-val myinv_f_f = thm "myinv_f_f";
-
-
-fun exh_thm_of (dt_info : datatype_info Symtab.table) tname =
- #exhaustion (the (Symtab.lookup dt_info tname));
-
-(******************************************************************************)
-
-fun representation_proofs (config : datatype_config) (dt_info : datatype_info Symtab.table)
- new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
- let
- val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
- val node_name = "Datatype.node";
- val In0_name = "Datatype.In0";
- val In1_name = "Datatype.In1";
- val Scons_name = "Datatype.Scons";
- val Leaf_name = "Datatype.Leaf";
- val Numb_name = "Datatype.Numb";
- val Lim_name = "Datatype.Lim";
- val Suml_name = "Datatype.Suml";
- val Sumr_name = "Datatype.Sumr";
-
- val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
- In0_eq, In1_eq, In0_not_In1, In1_not_In0,
- Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy)
- ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
- "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
- "Lim_inject", "Suml_inject", "Sumr_inject"];
-
- val descr' = flat descr;
-
- val big_name = space_implode "_" new_type_names;
- val thy1 = add_path (#flat_names config) big_name thy;
- val big_rec_name = big_name ^ "_rep_set";
- val rep_set_names' =
- (if length descr' = 1 then [big_rec_name] else
- (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
-
- val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
- val leafTs' = get_nonrec_types descr' sorts;
- val branchTs = get_branching_types descr' sorts;
- val branchT = if null branchTs then HOLogic.unitT
- else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
- val arities = get_arities descr' \ 0;
- val unneeded_vars = hd tyvars \\ List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
- val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
- val oldTs = Library.drop (length (hd descr), recTs);
- val sumT = if null leafTs then HOLogic.unitT
- else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
- val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
- val UnivT = HOLogic.mk_setT Univ_elT;
- val UnivT' = Univ_elT --> HOLogic.boolT;
- val Collect = Const ("Collect", UnivT' --> UnivT);
-
- val In0 = Const (In0_name, Univ_elT --> Univ_elT);
- val In1 = Const (In1_name, Univ_elT --> Univ_elT);
- val Leaf = Const (Leaf_name, sumT --> Univ_elT);
- val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
-
- (* make injections needed for embedding types in leaves *)
-
- fun mk_inj T' x =
- let
- fun mk_inj' T n i =
- if n = 1 then x else
- let val n2 = n div 2;
- val Type (_, [T1, T2]) = T
- in
- if i <= n2 then
- Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
- else
- Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
- end
- in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
- end;
-
- (* make injections for constructors *)
-
- fun mk_univ_inj ts = BalancedTree.access
- {left = fn t => In0 $ t,
- right = fn t => In1 $ t,
- init =
- if ts = [] then Const (@{const_name undefined}, Univ_elT)
- else foldr1 (HOLogic.mk_binop Scons_name) ts};
-
- (* function spaces *)
-
- fun mk_fun_inj T' x =
- let
- fun mk_inj T n i =
- if n = 1 then x else
- let
- val n2 = n div 2;
- val Type (_, [T1, T2]) = T;
- fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
- in
- if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
- else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
- end
- in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
- end;
-
- val mk_lim = List.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t)));
-
- (************** generate introduction rules for representing set **********)
-
- val _ = message config "Constructing representing sets ...";
-
- (* make introduction rule for a single constructor *)
-
- fun make_intr s n (i, (_, cargs)) =
- let
- fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of
- (dts, DtRec k) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) dts;
- val free_t =
- app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
- in (j + 1, list_all (map (pair "x") Ts,
- HOLogic.mk_Trueprop
- (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
- mk_lim free_t Ts :: ts)
- end
- | _ =>
- let val T = typ_of_dtyp descr' sorts dt
- in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
- end);
-
- val (_, prems, ts) = List.foldr mk_prem (1, [], []) cargs;
- val concl = HOLogic.mk_Trueprop
- (Free (s, UnivT') $ mk_univ_inj ts n i)
- in Logic.list_implies (prems, concl)
- end;
-
- val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
- map (make_intr rep_set_name (length constrs))
- ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
-
- val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
- (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
-
- (********************************* typedef ********************************)
-
- val (typedefs, thy3) = thy2 |>
- parent_path (#flat_names config) |>
- fold_map (fn ((((name, mx), tvs), c), name') =>
- TypedefPackage.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
- (Collect $ Const (c, UnivT')) NONE
- (rtac exI 1 THEN rtac CollectI 1 THEN
- QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)))
- (types_syntax ~~ tyvars ~~
- (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
- add_path (#flat_names config) big_name;
-
- (*********************** definition of constructors ***********************)
-
- val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
- val rep_names = map (curry op ^ "Rep_") new_type_names;
- val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
- (1 upto (length (flat (tl descr))));
- val all_rep_names = map (Sign.intern_const thy3) rep_names @
- map (Sign.full_bname thy3) rep_names';
-
- (* isomorphism declarations *)
-
- val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
- (oldTs ~~ rep_names');
-
- (* constructor definitions *)
-
- fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
- let
- fun constr_arg (dt, (j, l_args, r_args)) =
- let val T = typ_of_dtyp descr' sorts dt;
- val free_t = mk_Free "x" T j
- in (case (strip_dtyp dt, strip_type T) of
- ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
- (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
- app_bnds free_t (length Us)) Us :: r_args)
- | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
- end;
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
- val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
- val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
- val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
- val lhs = list_comb (Const (cname, constrT), l_args);
- val rhs = mk_univ_inj r_args n i;
- val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
- val def_name = Long_Name.base_name cname ^ "_def";
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
- val ([def_thm], thy') =
- thy
- |> Sign.add_consts_i [(cname', constrT, mx)]
- |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
-
- in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
-
- (* constructor definitions for datatype *)
-
- fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
- ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
- let
- val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
- val rep_const = cterm_of thy
- (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
- val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
- val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
- ((add_path (#flat_names config) tname thy, defs, [], 1), constrs ~~ constr_syntax)
- in
- (parent_path (#flat_names config) thy', defs', eqns @ [eqns'],
- rep_congs @ [cong'], dist_lemmas @ [dist])
- end;
-
- val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy3 |> Sign.add_consts_i iso_decls |> parent_path (#flat_names config), [], [], [], []),
- hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
-
- (*********** isomorphisms for new types (introduced by typedef) ***********)
-
- val _ = message config "Proving isomorphism properties ...";
-
- val newT_iso_axms = map (fn (_, td) =>
- (collect_simp (#Abs_inverse td), #Rep_inverse td,
- collect_simp (#Rep td))) typedefs;
-
- val newT_iso_inj_thms = map (fn (_, td) =>
- (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
-
- (********* isomorphisms between existing types and "unfolded" types *******)
-
- (*---------------------------------------------------------------------*)
- (* isomorphisms are defined using primrec-combinators: *)
- (* generate appropriate functions for instantiating primrec-combinator *)
- (* *)
- (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
- (* *)
- (* also generate characteristic equations for isomorphisms *)
- (* *)
- (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
- (*---------------------------------------------------------------------*)
-
- fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
- let
- val argTs = map (typ_of_dtyp descr' sorts) cargs;
- val T = List.nth (recTs, k);
- val rep_name = List.nth (all_rep_names, k);
- val rep_const = Const (rep_name, T --> Univ_elT);
- val constr = Const (cname, argTs ---> T);
-
- fun process_arg ks' ((i2, i2', ts, Ts), dt) =
- let
- val T' = typ_of_dtyp descr' sorts dt;
- val (Us, U) = strip_type T'
- in (case strip_dtyp dt of
- (_, DtRec j) => if j mem ks' then
- (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
- (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
- Ts @ [Us ---> Univ_elT])
- else
- (i2 + 1, i2', ts @ [mk_lim
- (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
- app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
- | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
- end;
-
- val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
- val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
- val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
- val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
-
- val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
-
- in (fs @ [f], eqns @ [eqn], i + 1) end;
-
- (* define isomorphisms for all mutually recursive datatypes in list ds *)
-
- fun make_iso_defs (ds, (thy, char_thms)) =
- let
- val ks = map fst ds;
- val (_, (tname, _, _)) = hd ds;
- val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
-
- fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
- let
- val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
- ((fs, eqns, 1), constrs);
- val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
- in (fs', eqns', isos @ [iso]) end;
-
- val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
- val fTs = map fastype_of fs;
- val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
- Logic.mk_equals (Const (iso_name, T --> Univ_elT),
- list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
- val (def_thms, thy') =
- apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy);
-
- (* prove characteristic equations *)
-
- val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
- val char_thms' = map (fn eqn => SkipProof.prove_global thy' [] [] eqn
- (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
-
- in (thy', char_thms' @ char_thms) end;
-
- val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs
- (add_path (#flat_names config) big_name thy4, []) (tl descr));
-
- (* prove isomorphism properties *)
-
- fun mk_funs_inv thy thm =
- let
- val prop = Thm.prop_of thm;
- val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
- (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
- val used = OldTerm.add_term_tfree_names (a, []);
-
- fun mk_thm i =
- let
- val Ts = map (TFree o rpair HOLogic.typeS)
- (Name.variant_list used (replicate i "'t"));
- val f = Free ("f", Ts ---> U)
- in SkipProof.prove_global thy [] [] (Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.list_all
- (map (pair "x") Ts, S $ app_bnds f i)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
- r $ (a $ app_bnds f i)), f))))
- (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
- REPEAT (etac allE 1), rtac thm 1, atac 1])
- end
- in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
-
- (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *)
-
- val fun_congs = map (fn T => make_elim (Drule.instantiate'
- [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
-
- fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
- let
- val (_, (tname, _, _)) = hd ds;
- val {induction, ...} = the (Symtab.lookup dt_info tname);
-
- fun mk_ind_concl (i, _) =
- let
- val T = List.nth (recTs, i);
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
- val rep_set_name = List.nth (rep_set_names, i)
- in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
- HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
- HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
- Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
- end;
-
- val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
-
- val rewrites = map mk_meta_eq iso_char_thms;
- val inj_thms' = map snd newT_iso_inj_thms @
- map (fn r => r RS @{thm injD}) inj_thms;
-
- val inj_thm = SkipProof.prove_global thy5 [] []
- (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
- [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- REPEAT (EVERY
- [rtac allI 1, rtac impI 1,
- exh_tac (exh_thm_of dt_info) 1,
- REPEAT (EVERY
- [hyp_subst_tac 1,
- rewrite_goals_tac rewrites,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
- ORELSE (EVERY
- [REPEAT (eresolve_tac (Scons_inject ::
- map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
- REPEAT (cong_tac 1), rtac refl 1,
- REPEAT (atac 1 ORELSE (EVERY
- [REPEAT (rtac ext 1),
- REPEAT (eresolve_tac (mp :: allE ::
- map make_elim (Suml_inject :: Sumr_inject ::
- Lim_inject :: inj_thms') @ fun_congs) 1),
- atac 1]))])])])]);
-
- val inj_thms'' = map (fn r => r RS @{thm datatype_injI})
- (split_conj_thm inj_thm);
-
- val elem_thm =
- SkipProof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
- (fn _ =>
- EVERY [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- rewrite_goals_tac rewrites,
- REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
- ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
-
- in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
- end;
-
- val (iso_inj_thms_unfolded, iso_elem_thms) = List.foldr prove_iso_thms
- ([], map #3 newT_iso_axms) (tl descr);
- val iso_inj_thms = map snd newT_iso_inj_thms @
- map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
-
- (* prove dt_rep_set_i x --> x : range dt_Rep_i *)
-
- fun mk_iso_t (((set_name, iso_name), i), T) =
- let val isoT = T --> Univ_elT
- in HOLogic.imp $
- (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
- (if i < length newTs then HOLogic.true_const
- else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
- Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
- Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
- end;
-
- val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
- (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
-
- (* all the theorems are proved by one single simultaneous induction *)
-
- val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq}))
- iso_inj_thms_unfolded;
-
- val iso_thms = if length descr = 1 then [] else
- Library.drop (length newTs, split_conj_thm
- (SkipProof.prove_global thy5 [] [] iso_t (fn _ => EVERY
- [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- REPEAT (rtac TrueI 1),
- rewrite_goals_tac (mk_meta_eq choice_eq ::
- symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
- rewrite_goals_tac (map symmetric range_eqs),
- REPEAT (EVERY
- [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
- maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
- TRY (hyp_subst_tac 1),
- rtac (sym RS range_eqI) 1,
- resolve_tac iso_char_thms 1])])));
-
- val Abs_inverse_thms' =
- map #1 newT_iso_axms @
- map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
- iso_inj_thms_unfolded iso_thms;
-
- val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
-
- (******************* freeness theorems for constructors *******************)
-
- val _ = message config "Proving freeness of constructors ...";
-
- (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *)
-
- fun prove_constr_rep_thm eqn =
- let
- val inj_thms = map fst newT_iso_inj_thms;
- val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
- in SkipProof.prove_global thy5 [] [] eqn (fn _ => EVERY
- [resolve_tac inj_thms 1,
- rewrite_goals_tac rewrites,
- rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac iso_elem_thms 1)])
- end;
-
- (*--------------------------------------------------------------*)
- (* constr_rep_thms and rep_congs are used to prove distinctness *)
- (* of constructors. *)
- (*--------------------------------------------------------------*)
-
- val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
-
- val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
- dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
- (constr_rep_thms ~~ dist_lemmas);
-
- fun prove_distinct_thms _ _ (_, []) = []
- | prove_distinct_thms lim dist_rewrites' (k, ts as _ :: _) =
- if k >= lim then [] else let
- (*number of constructors < distinctness_limit : C_i ... ~= C_j ...*)
- fun prove [] = []
- | prove (t :: ts) =
- let
- val dist_thm = SkipProof.prove_global thy5 [] [] t (fn _ =>
- EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
- in dist_thm :: standard (dist_thm RS not_sym) :: prove ts end;
- in prove ts end;
-
- val distinct_thms = DatatypeProp.make_distincts descr sorts
- |> map2 (prove_distinct_thms
- (Config.get_thy thy5 distinctness_limit)) dist_rewrites;
-
- val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
- if length constrs < Config.get_thy thy5 distinctness_limit
- then FewConstrs dists
- else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
- constr_rep_thms ~~ rep_congs ~~ distinct_thms);
-
- (* prove injectivity of constructors *)
-
- fun prove_constr_inj_thm rep_thms t =
- let val inj_thms = Scons_inject :: (map make_elim
- (iso_inj_thms @
- [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
- Lim_inject, Suml_inject, Sumr_inject]))
- in SkipProof.prove_global thy5 [] [] t (fn _ => EVERY
- [rtac iffI 1,
- REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
- dresolve_tac rep_congs 1, dtac box_equals 1,
- REPEAT (resolve_tac rep_thms 1),
- REPEAT (eresolve_tac inj_thms 1),
- REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
- REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
- atac 1]))])
- end;
-
- val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
- ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
-
- val ((constr_inject', distinct_thms'), thy6) =
- thy5
- |> parent_path (#flat_names config)
- |> store_thmss "inject" new_type_names constr_inject
- ||>> store_thmss "distinct" new_type_names distinct_thms;
-
- (*************************** induction theorem ****************************)
-
- val _ = message config "Proving induction rule for datatypes ...";
-
- val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
- (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
- val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
-
- fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
- let
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
- mk_Free "x" T i;
-
- val Abs_t = if i < length newTs then
- Const (Sign.intern_const thy6
- ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
- else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
- Const (List.nth (all_rep_names, i), T --> Univ_elT)
-
- in (prems @ [HOLogic.imp $
- (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
- (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
- concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
- end;
-
- val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
-
- val cert = cterm_of thy6;
-
- val indrule_lemma = SkipProof.prove_global thy6 [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
- [REPEAT (etac conjE 1),
- REPEAT (EVERY
- [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
- etac mp 1, resolve_tac iso_elem_thms 1])]);
-
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
- val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
- map (Free o apfst fst o dest_Var) Ps;
- val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
-
- val dt_induct_prop = DatatypeProp.make_ind descr sorts;
- val dt_induct = SkipProof.prove_global thy6 []
- (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
- (fn {prems, ...} => EVERY
- [rtac indrule_lemma' 1,
- (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms 1),
- simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
- DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
- (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
-
- val ([dt_induct'], thy7) =
- thy6
- |> Sign.add_path big_name
- |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
- ||> Sign.parent_path
- ||> Theory.checkpoint;
-
- in
- ((constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct'), thy7)
- end;
-
-end;
--- a/src/HOL/Tools/function_package/auto_term.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-(* Title: HOL/Tools/function_package/auto_term.ML
- ID: $Id$
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Method "relation" to commence a termination proof using a user-specified relation.
-*)
-
-signature FUNDEF_RELATION =
-sig
- val relation_tac: Proof.context -> term -> int -> tactic
- val setup: theory -> theory
-end
-
-structure FundefRelation : FUNDEF_RELATION =
-struct
-
-fun inst_thm ctxt rel st =
- let
- val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
- val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
- val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
- val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
- in
- Drule.cterm_instantiate [(Rvar, rel')] st'
- end
-
-fun relation_tac ctxt rel i =
- TRY (FundefCommon.apply_termination_rule ctxt i)
- THEN PRIMITIVE (inst_thm ctxt rel)
-
-val setup =
- Method.setup @{binding relation}
- (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
- "proves termination using a user-specified wellfounded relation"
-
-end
--- a/src/HOL/Tools/function_package/context_tree.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-(* Title: HOL/Tools/function_package/context_tree.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Builds and traverses trees of nested contexts along a term.
-*)
-
-signature FUNDEF_CTXTREE =
-sig
- type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
- type ctx_tree
-
- (* FIXME: This interface is a mess and needs to be cleaned up! *)
- val get_fundef_congs : Proof.context -> thm list
- val add_fundef_cong : thm -> Context.generic -> Context.generic
- val map_fundef_congs : (thm list -> thm list) -> Context.generic -> Context.generic
-
- val cong_add: attribute
- val cong_del: attribute
-
- val mk_tree: (string * typ) -> term -> Proof.context -> term -> ctx_tree
-
- val inst_tree: theory -> term -> term -> ctx_tree -> ctx_tree
-
- val export_term : ctxt -> term -> term
- val export_thm : theory -> ctxt -> thm -> thm
- val import_thm : theory -> ctxt -> thm -> thm
-
- val traverse_tree :
- (ctxt -> term ->
- (ctxt * thm) list ->
- (ctxt * thm) list * 'b ->
- (ctxt * thm) list * 'b)
- -> ctx_tree -> 'b -> 'b
-
- val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list
-end
-
-structure FundefCtxTree : FUNDEF_CTXTREE =
-struct
-
-type ctxt = (string * typ) list * thm list
-
-open FundefCommon
-open FundefLib
-
-structure FundefCongs = GenericDataFun
-(
- type T = thm list
- val empty = []
- val extend = I
- fun merge _ = Thm.merge_thms
-);
-
-val get_fundef_congs = FundefCongs.get o Context.Proof
-val map_fundef_congs = FundefCongs.map
-val add_fundef_cong = FundefCongs.map o Thm.add_thm
-
-(* congruence rules *)
-
-val cong_add = Thm.declaration_attribute (map_fundef_congs o Thm.add_thm o safe_mk_meta_eq);
-val cong_del = Thm.declaration_attribute (map_fundef_congs o Thm.del_thm o safe_mk_meta_eq);
-
-
-type depgraph = int IntGraph.T
-
-datatype ctx_tree
- = Leaf of term
- | Cong of (thm * depgraph * (ctxt * ctx_tree) list)
- | RCall of (term * ctx_tree)
-
-
-(* Maps "Trueprop A = B" to "A" *)
-val rhs_of = snd o HOLogic.dest_eq o HOLogic.dest_Trueprop
-
-
-(*** Dependency analysis for congruence rules ***)
-
-fun branch_vars t =
- let
- val t' = snd (dest_all_all t)
- val (assumes, concl) = Logic.strip_horn t'
- in (fold Term.add_vars assumes [], Term.add_vars concl [])
- end
-
-fun cong_deps crule =
- let
- val num_branches = map_index (apsnd branch_vars) (prems_of crule)
- in
- IntGraph.empty
- |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches
- |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) =>
- if i = j orelse null (c1 inter t2)
- then I else IntGraph.add_edge_acyclic (i,j))
- num_branches num_branches
- end
-
-val default_congs = map (fn c => c RS eq_reflection) [@{thm "cong"}, @{thm "ext"}]
-
-
-
-(* Called on the INSTANTIATED branches of the congruence rule *)
-fun mk_branch ctx t =
- let
- val (ctx', fixes, impl) = dest_all_all_ctx ctx t
- val (assms, concl) = Logic.strip_horn impl
- in
- (ctx', fixes, assms, rhs_of concl)
- end
-
-fun find_cong_rule ctx fvar h ((r,dep)::rs) t =
- (let
- val thy = ProofContext.theory_of ctx
-
- val tt' = Logic.mk_equals (Pattern.rewrite_term thy [(Free fvar, h)] [] t, t)
- val (c, subs) = (concl_of r, prems_of r)
-
- val subst = Pattern.match (ProofContext.theory_of ctx) (c, tt') (Vartab.empty, Vartab.empty)
- val branches = map (mk_branch ctx o Envir.beta_norm o Envir.subst_vars subst) subs
- val inst = map (fn v => (cterm_of thy (Var v), cterm_of thy (Envir.subst_vars subst (Var v)))) (Term.add_vars c [])
- in
- (cterm_instantiate inst r, dep, branches)
- end
- handle Pattern.MATCH => find_cong_rule ctx fvar h rs t)
- | find_cong_rule _ _ _ [] _ = sys_error "function_package/context_tree.ML: No cong rule found!"
-
-
-fun mk_tree fvar h ctxt t =
- let
- val congs = get_fundef_congs ctxt
- val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) (* FIXME: Save in theory *)
-
- fun matchcall (a $ b) = if a = Free fvar then SOME b else NONE
- | matchcall _ = NONE
-
- fun mk_tree' ctx t =
- case matchcall t of
- SOME arg => RCall (t, mk_tree' ctx arg)
- | NONE =>
- if not (exists_subterm (fn Free v => v = fvar | _ => false) t) then Leaf t
- else
- let val (r, dep, branches) = find_cong_rule ctx fvar h congs_deps t in
- Cong (r, dep,
- map (fn (ctx', fixes, assumes, st) =>
- ((fixes, map (assume o cterm_of (ProofContext.theory_of ctx)) assumes),
- mk_tree' ctx' st)) branches)
- end
- in
- mk_tree' ctxt t
- end
-
-
-fun inst_tree thy fvar f tr =
- let
- val cfvar = cterm_of thy fvar
- val cf = cterm_of thy f
-
- fun inst_term t =
- subst_bound(f, abstract_over (fvar, t))
-
- val inst_thm = forall_elim cf o forall_intr cfvar
-
- fun inst_tree_aux (Leaf t) = Leaf t
- | inst_tree_aux (Cong (crule, deps, branches)) =
- Cong (inst_thm crule, deps, map inst_branch branches)
- | inst_tree_aux (RCall (t, str)) =
- RCall (inst_term t, inst_tree_aux str)
- and inst_branch ((fxs, assms), str) =
- ((fxs, map (assume o cterm_of thy o inst_term o prop_of) assms), inst_tree_aux str)
- in
- inst_tree_aux tr
- end
-
-
-(* Poor man's contexts: Only fixes and assumes *)
-fun compose (fs1, as1) (fs2, as2) = (fs1 @ fs2, as1 @ as2)
-
-fun export_term (fixes, assumes) =
- fold_rev (curry Logic.mk_implies o prop_of) assumes
- #> fold_rev (Logic.all o Free) fixes
-
-fun export_thm thy (fixes, assumes) =
- fold_rev (implies_intr o cprop_of) assumes
- #> fold_rev (forall_intr o cterm_of thy o Free) fixes
-
-fun import_thm thy (fixes, athms) =
- fold (forall_elim o cterm_of thy o Free) fixes
- #> fold Thm.elim_implies athms
-
-
-(* folds in the order of the dependencies of a graph. *)
-fun fold_deps G f x =
- let
- fun fill_table i (T, x) =
- case Inttab.lookup T i of
- SOME _ => (T, x)
- | NONE =>
- let
- val (T', x') = fold fill_table (IntGraph.imm_succs G i) (T, x)
- val (v, x'') = f (the o Inttab.lookup T') i x'
- in
- (Inttab.update (i, v) T', x'')
- end
-
- val (T, x) = fold fill_table (IntGraph.keys G) (Inttab.empty, x)
- in
- (Inttab.fold (cons o snd) T [], x)
- end
-
-fun traverse_tree rcOp tr =
- let
- fun traverse_help ctx (Leaf _) _ x = ([], x)
- | traverse_help ctx (RCall (t, st)) u x =
- rcOp ctx t u (traverse_help ctx st u x)
- | traverse_help ctx (Cong (_, deps, branches)) u x =
- let
- fun sub_step lu i x =
- let
- val (ctx', subtree) = nth branches i
- val used = fold_rev (append o lu) (IntGraph.imm_succs deps i) u
- val (subs, x') = traverse_help (compose ctx ctx') subtree used x
- val exported_subs = map (apfst (compose ctx')) subs (* FIXME: Right order of composition? *)
- in
- (exported_subs, x')
- end
- in
- fold_deps deps sub_step x
- |> apfst flat
- end
- in
- snd o traverse_help ([], []) tr []
- end
-
-fun rewrite_by_tree thy h ih x tr =
- let
- fun rewrite_help _ _ x (Leaf t) = (reflexive (cterm_of thy t), x)
- | rewrite_help fix h_as x (RCall (_ $ arg, st)) =
- let
- val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *)
-
- val iha = import_thm thy (fix, h_as) ha (* (a', h a') : G *)
- |> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner))))
- (* (a, h a) : G *)
- val inst_ih = instantiate' [] [SOME (cterm_of thy arg)] ih
- val eq = implies_elim (implies_elim inst_ih lRi) iha (* h a = f a *)
-
- val h_a'_eq_h_a = combination (reflexive (cterm_of thy h)) inner
- val h_a_eq_f_a = eq RS eq_reflection
- val result = transitive h_a'_eq_h_a h_a_eq_f_a
- in
- (result, x')
- end
- | rewrite_help fix h_as x (Cong (crule, deps, branches)) =
- let
- fun sub_step lu i x =
- let
- val ((fixes, assumes), st) = nth branches i
- val used = map lu (IntGraph.imm_succs deps i)
- |> map (fn u_eq => (u_eq RS sym) RS eq_reflection)
- |> filter_out Thm.is_reflexive
-
- val assumes' = map (simplify (HOL_basic_ss addsimps used)) assumes
-
- val (subeq, x') = rewrite_help (fix @ fixes) (h_as @ assumes') x st
- val subeq_exp = export_thm thy (fixes, assumes) (subeq RS meta_eq_to_obj_eq)
- in
- (subeq_exp, x')
- end
-
- val (subthms, x') = fold_deps deps sub_step x
- in
- (fold_rev (curry op COMP) subthms crule, x')
- end
- in
- rewrite_help [] [] x tr
- end
-
-end
--- a/src/HOL/Tools/function_package/decompose.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-(* Title: HOL/Tools/function_package/decompose.ML
- Author: Alexander Krauss, TU Muenchen
-
-Graph decomposition using "Shallow Dependency Pairs".
-*)
-
-signature DECOMPOSE =
-sig
-
- val derive_chains : Proof.context -> tactic
- -> (Termination.data -> int -> tactic)
- -> Termination.data -> int -> tactic
-
- val decompose_tac : Proof.context -> tactic
- -> Termination.ttac
-
-end
-
-structure Decompose : DECOMPOSE =
-struct
-
-structure TermGraph = GraphFun(type key = term val ord = TermOrd.fast_term_ord);
-
-
-fun derive_chains ctxt chain_tac cont D = Termination.CALLS (fn (cs, i) =>
- let
- val thy = ProofContext.theory_of ctxt
-
- fun prove_chain c1 c2 D =
- if is_some (Termination.get_chain D c1 c2) then D else
- let
- val goal = HOLogic.mk_eq (HOLogic.mk_binop @{const_name "Relation.rel_comp"} (c1, c2),
- Const (@{const_name Set.empty}, fastype_of c1))
- |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
-
- val chain = case FundefLib.try_proof (cterm_of thy goal) chain_tac of
- FundefLib.Solved thm => SOME thm
- | _ => NONE
- in
- Termination.note_chain c1 c2 chain D
- end
- in
- cont (fold_product prove_chain cs cs D) i
- end)
-
-
-fun mk_dgraph D cs =
- TermGraph.empty
- |> fold (fn c => TermGraph.new_node (c,())) cs
- |> fold_product (fn c1 => fn c2 =>
- if is_none (Termination.get_chain D c1 c2 |> the_default NONE)
- then TermGraph.add_edge (c1, c2) else I)
- cs cs
-
-
-fun ucomp_empty_tac T =
- REPEAT_ALL_NEW (rtac @{thm union_comp_emptyR}
- ORELSE' rtac @{thm union_comp_emptyL}
- ORELSE' SUBGOAL (fn (_ $ (_ $ (_ $ c1 $ c2) $ _), i) => rtac (T c1 c2) i))
-
-fun regroup_calls_tac cs = Termination.CALLS (fn (cs', i) =>
- let
- val is = map (fn c => find_index (curry op aconv c) cs') cs
- in
- CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv is))) i
- end)
-
-
-fun solve_trivial_tac D = Termination.CALLS
-(fn ([c], i) =>
- (case Termination.get_chain D c c of
- SOME (SOME thm) => rtac @{thm wf_no_loop} i
- THEN rtac thm i
- | _ => no_tac)
- | _ => no_tac)
-
-fun decompose_tac' ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
- let
- val G = mk_dgraph D cs
- val sccs = TermGraph.strong_conn G
-
- fun split [SCC] i = (solve_trivial_tac D i ORELSE cont D i)
- | split (SCC::rest) i =
- regroup_calls_tac SCC i
- THEN rtac @{thm wf_union_compatible} i
- THEN rtac @{thm less_by_empty} (i + 2)
- THEN ucomp_empty_tac (the o the oo Termination.get_chain D) (i + 2)
- THEN split rest (i + 1)
- THEN (solve_trivial_tac D i ORELSE cont D i)
- in
- if length sccs > 1 then split sccs i
- else solve_trivial_tac D i ORELSE err_cont D i
- end)
-
-fun decompose_tac ctxt chain_tac cont err_cont =
- derive_chains ctxt chain_tac
- (decompose_tac' ctxt cont err_cont)
-
-fun auto_decompose_tac ctxt =
- Termination.TERMINATION ctxt
- (decompose_tac ctxt (auto_tac (local_clasimpset_of ctxt))
- (K (K all_tac)) (K (K no_tac)))
-
-
-end
--- a/src/HOL/Tools/function_package/descent.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,44 +0,0 @@
-(* Title: HOL/Tools/function_package/descent.ML
- Author: Alexander Krauss, TU Muenchen
-
-Descent proofs for termination
-*)
-
-
-signature DESCENT =
-sig
-
- val derive_diag : Proof.context -> tactic -> (Termination.data -> int -> tactic)
- -> Termination.data -> int -> tactic
-
- val derive_all : Proof.context -> tactic -> (Termination.data -> int -> tactic)
- -> Termination.data -> int -> tactic
-
-end
-
-
-structure Descent : DESCENT =
-struct
-
-fun gen_descent diag ctxt tac cont D = Termination.CALLS (fn (cs, i) =>
- let
- val thy = ProofContext.theory_of ctxt
- val measures_of = Termination.get_measures D
-
- fun derive c D =
- let
- val (_, p, _, q, _, _) = Termination.dest_call D c
- in
- if diag andalso p = q
- then fold (fn m => Termination.derive_descent thy tac c m m) (measures_of p) D
- else fold_product (Termination.derive_descent thy tac c)
- (measures_of p) (measures_of q) D
- end
- in
- cont (FundefCommon.PROFILE "deriving descents" (fold derive cs) D) i
- end)
-
-val derive_diag = gen_descent true
-val derive_all = gen_descent false
-
-end
--- a/src/HOL/Tools/function_package/fundef_common.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,343 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_common.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Common definitions and other infrastructure.
-*)
-
-structure FundefCommon =
-struct
-
-local open FundefLib in
-
-(* Profiling *)
-val profile = ref false;
-
-fun PROFILE msg = if !profile then timeap_msg msg else I
-
-
-val acc_const_name = @{const_name "accp"}
-fun mk_acc domT R =
- Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R
-
-val function_name = suffix "C"
-val graph_name = suffix "_graph"
-val rel_name = suffix "_rel"
-val dom_name = suffix "_dom"
-
-(* Termination rules *)
-
-structure TerminationRule = GenericDataFun
-(
- type T = thm list
- val empty = []
- val extend = I
- fun merge _ = Thm.merge_thms
-);
-
-val get_termination_rules = TerminationRule.get
-val store_termination_rule = TerminationRule.map o cons
-val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
-
-
-(* Function definition result data *)
-
-datatype fundef_result =
- FundefResult of
- {
- fs: term list,
- G: term,
- R: term,
-
- psimps : thm list,
- trsimps : thm list option,
-
- simple_pinducts : thm list,
- cases : thm,
- termination : thm,
- domintros : thm list option
- }
-
-
-datatype fundef_context_data =
- FundefCtxData of
- {
- defname : string,
-
- (* contains no logical entities: invariant under morphisms *)
- add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list
- -> local_theory -> thm list * local_theory,
- case_names : string list,
-
- fs : term list,
- R : term,
-
- psimps: thm list,
- pinducts: thm list,
- termination: thm
- }
-
-fun morph_fundef_data (FundefCtxData {add_simps, case_names, fs, R,
- psimps, pinducts, termination, defname}) phi =
- let
- val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
- val name = Binding.name_of o Morphism.binding phi o Binding.name
- in
- FundefCtxData { add_simps = add_simps, case_names = case_names,
- fs = map term fs, R = term R, psimps = fact psimps,
- pinducts = fact pinducts, termination = thm termination,
- defname = name defname }
- end
-
-structure FundefData = GenericDataFun
-(
- type T = (term * fundef_context_data) Item_Net.T;
- val empty = Item_Net.init
- (op aconv o pairself fst : (term * fundef_context_data) * (term * fundef_context_data) -> bool)
- fst;
- val copy = I;
- val extend = I;
- fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
-);
-
-val get_fundef = FundefData.get o Context.Proof;
-
-
-(* Generally useful?? *)
-fun lift_morphism thy f =
- let
- val term = Drule.term_rule thy f
- in
- Morphism.thm_morphism f $> Morphism.term_morphism term
- $> Morphism.typ_morphism (Logic.type_map term)
- end
-
-fun import_fundef_data t ctxt =
- let
- val thy = ProofContext.theory_of ctxt
- val ct = cterm_of thy t
- val inst_morph = lift_morphism thy o Thm.instantiate
-
- fun match (trm, data) =
- SOME (morph_fundef_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
- handle Pattern.MATCH => NONE
- in
- get_first match (Item_Net.retrieve (get_fundef ctxt) t)
- end
-
-fun import_last_fundef ctxt =
- case Item_Net.content (get_fundef ctxt) of
- [] => NONE
- | (t, data) :: _ =>
- let
- val ([t'], ctxt') = Variable.import_terms true [t] ctxt
- in
- import_fundef_data t' ctxt'
- end
-
-val all_fundef_data = Item_Net.content o get_fundef
-
-fun add_fundef_data (data as FundefCtxData {fs, termination, ...}) =
- FundefData.map (fold (fn f => Item_Net.insert (f, data)) fs)
- #> store_termination_rule termination
-
-
-(* Simp rules for termination proofs *)
-
-structure TerminationSimps = NamedThmsFun
-(
- val name = "termination_simp"
- val description = "Simplification rule for termination proofs"
-);
-
-
-(* Default Termination Prover *)
-
-structure TerminationProver = GenericDataFun
-(
- type T = Proof.context -> Proof.method
- val empty = (fn _ => error "Termination prover not configured")
- val extend = I
- fun merge _ (a,b) = b (* FIXME *)
-);
-
-val set_termination_prover = TerminationProver.put
-val get_termination_prover = TerminationProver.get o Context.Proof
-
-
-(* Configuration management *)
-datatype fundef_opt
- = Sequential
- | Default of string
- | DomIntros
- | Tailrec
-
-datatype fundef_config
- = FundefConfig of
- {
- sequential: bool,
- default: string,
- domintros: bool,
- tailrec: bool
- }
-
-fun apply_opt Sequential (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=true, default=default, domintros=domintros, tailrec=tailrec}
- | apply_opt (Default d) (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=d, domintros=domintros, tailrec=tailrec}
- | apply_opt DomIntros (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=default, domintros=true,tailrec=tailrec}
- | apply_opt Tailrec (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=default, domintros=domintros,tailrec=true}
-
-val default_config =
- FundefConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*),
- domintros=false, tailrec=false }
-
-
-(* Analyzing function equations *)
-
-fun split_def ctxt geq =
- let
- fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
- val qs = Term.strip_qnt_vars "all" geq
- val imp = Term.strip_qnt_body "all" geq
- val (gs, eq) = Logic.strip_horn imp
-
- val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
- handle TERM _ => error (input_error "Not an equation")
-
- val (head, args) = strip_comb f_args
-
- val fname = fst (dest_Free head)
- handle TERM _ => error (input_error "Head symbol must not be a bound variable")
- in
- (fname, qs, gs, args, rhs)
- end
-
-(* Check for all sorts of errors in the input *)
-fun check_defs ctxt fixes eqs =
- let
- val fnames = map (fst o fst) fixes
-
- fun check geq =
- let
- fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
-
- val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
-
- val _ = fname mem fnames
- orelse input_error
- ("Head symbol of left hand side must be "
- ^ plural "" "one out of " fnames ^ commas_quote fnames)
-
- val _ = length args > 0 orelse input_error "Function has no arguments:"
-
- fun add_bvs t is = add_loose_bnos (t, 0, is)
- val rvs = (add_bvs rhs [] \\ fold add_bvs args [])
- |> map (fst o nth (rev qs))
-
- val _ = null rvs orelse input_error
- ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
- ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
-
- val _ = forall (not o Term.exists_subterm
- (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
- orelse input_error "Defined function may not occur in premises or arguments"
-
- val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
- val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
- val _ = null funvars
- orelse (warning (cat_lines
- ["Bound variable" ^ plural " " "s " funvars
- ^ commas_quote (map fst funvars) ^
- " occur" ^ plural "s" "" funvars ^ " in function position.",
- "Misspelled constructor???"]); true)
- in
- (fname, length args)
- end
-
- val _ = AList.group (op =) (map check eqs)
- |> map (fn (fname, ars) =>
- length (distinct (op =) ars) = 1
- orelse error ("Function " ^ quote fname ^
- " has different numbers of arguments in different equations"))
-
- fun check_sorts ((fname, fT), _) =
- Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
- orelse error (cat_lines
- ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
- setmp show_sorts true (Syntax.string_of_typ ctxt) fT])
-
- val _ = map check_sorts fixes
- in
- ()
- end
-
-(* Preprocessors *)
-
-type fixes = ((string * typ) * mixfix) list
-type 'a spec = (Attrib.binding * 'a list) list
-type preproc = fundef_config -> Proof.context -> fixes -> term spec
- -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
-
-val fname_of = fst o dest_Free o fst o strip_comb o fst
- o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
-
-fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
- | mk_case_names _ n 0 = []
- | mk_case_names _ n 1 = [n]
- | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
-
-fun empty_preproc check _ ctxt fixes spec =
- let
- val (bnds, tss) = split_list spec
- val ts = flat tss
- val _ = check ctxt fixes ts
- val fnames = map (fst o fst) fixes
- val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
-
- fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1)
- (indices ~~ xs)
- |> map (map snd)
-
- (* using theorem names for case name currently disabled *)
- val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
- in
- (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
- end
-
-structure Preprocessor = GenericDataFun
-(
- type T = preproc
- val empty : T = empty_preproc check_defs
- val extend = I
- fun merge _ (a, _) = a
-);
-
-val get_preproc = Preprocessor.get o Context.Proof
-val set_preproc = Preprocessor.map o K
-
-
-
-local
- structure P = OuterParse and K = OuterKeyword
-
- val option_parser =
- P.group "option" ((P.reserved "sequential" >> K Sequential)
- || ((P.reserved "default" |-- P.term) >> Default)
- || (P.reserved "domintros" >> K DomIntros)
- || (P.reserved "tailrec" >> K Tailrec))
-
- fun config_parser default =
- (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
- >> (fn opts => fold apply_opt opts default)
-in
- fun fundef_parser default_cfg =
- config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
-end
-
-
-end
-end
-
--- a/src/HOL/Tools/function_package/fundef_core.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,954 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_core.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions:
-Main functionality.
-*)
-
-signature FUNDEF_CORE =
-sig
- val prepare_fundef : FundefCommon.fundef_config
- -> string (* defname *)
- -> ((bstring * typ) * mixfix) list (* defined symbol *)
- -> ((bstring * typ) list * term list * term * term) list (* specification *)
- -> local_theory
-
- -> (term (* f *)
- * thm (* goalstate *)
- * (thm -> FundefCommon.fundef_result) (* continuation *)
- ) * local_theory
-
-end
-
-structure FundefCore : FUNDEF_CORE =
-struct
-
-val boolT = HOLogic.boolT
-val mk_eq = HOLogic.mk_eq
-
-open FundefLib
-open FundefCommon
-
-datatype globals =
- Globals of {
- fvar: term,
- domT: typ,
- ranT: typ,
- h: term,
- y: term,
- x: term,
- z: term,
- a: term,
- P: term,
- D: term,
- Pbool:term
-}
-
-
-datatype rec_call_info =
- RCInfo of
- {
- RIvs: (string * typ) list, (* Call context: fixes and assumes *)
- CCas: thm list,
- rcarg: term, (* The recursive argument *)
-
- llRI: thm,
- h_assum: term
- }
-
-
-datatype clause_context =
- ClauseContext of
- {
- ctxt : Proof.context,
-
- qs : term list,
- gs : term list,
- lhs: term,
- rhs: term,
-
- cqs: cterm list,
- ags: thm list,
- case_hyp : thm
- }
-
-
-fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
- ClauseContext { ctxt = ProofContext.transfer thy ctxt,
- qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
-
-
-datatype clause_info =
- ClauseInfo of
- {
- no: int,
- qglr : ((string * typ) list * term list * term * term),
- cdata : clause_context,
-
- tree: FundefCtxTree.ctx_tree,
- lGI: thm,
- RCs: rec_call_info list
- }
-
-
-(* Theory dependencies. *)
-val Pair_inject = @{thm Product_Type.Pair_inject};
-
-val acc_induct_rule = @{thm accp_induct_rule};
-
-val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
-val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
-val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
-
-val acc_downward = @{thm accp_downward};
-val accI = @{thm accp.accI};
-val case_split = @{thm HOL.case_split};
-val fundef_default_value = @{thm FunDef.fundef_default_value};
-val not_acc_down = @{thm not_accp_down};
-
-
-
-fun find_calls tree =
- let
- fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
- | add_Ri _ _ _ _ = raise Match
- in
- rev (FundefCtxTree.traverse_tree add_Ri tree [])
- end
-
-
-(** building proof obligations *)
-
-fun mk_compat_proof_obligations domT ranT fvar f glrs =
- let
- fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
- let
- val shift = incr_boundvars (length qs')
- in
- Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
- HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
- |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
- |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
- |> curry abstract_over fvar
- |> curry subst_bound f
- end
- in
- map mk_impl (unordered_pairs glrs)
- end
-
-
-fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
- let
- fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
- HOLogic.mk_Trueprop Pbool
- |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- in
- HOLogic.mk_Trueprop Pbool
- |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
- |> mk_forall_rename ("x", x)
- |> mk_forall_rename ("P", Pbool)
- end
-
-(** making a context with it's own local bindings **)
-
-fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
- let
- val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
- |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
-
- val thy = ProofContext.theory_of ctxt'
-
- fun inst t = subst_bounds (rev qs, t)
- val gs = map inst pre_gs
- val lhs = inst pre_lhs
- val rhs = inst pre_rhs
-
- val cqs = map (cterm_of thy) qs
- val ags = map (assume o cterm_of thy) gs
-
- val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
- in
- ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
- cqs = cqs, ags = ags, case_hyp = case_hyp }
- end
-
-
-(* lowlevel term function *)
-fun abstract_over_list vs body =
- let
- exception SAME;
- fun abs lev v tm =
- if v aconv tm then Bound lev
- else
- (case tm of
- Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
- | t $ u => (abs lev v t $ (abs lev v u handle SAME => u) handle SAME => t $ abs lev v u)
- | _ => raise SAME);
- in
- fold_index (fn (i,v) => fn t => abs i v t handle SAME => t) vs body
- end
-
-
-
-fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
- let
- val Globals {h, fvar, x, ...} = globals
-
- val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
- val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
-
- (* Instantiate the GIntro thm with "f" and import into the clause context. *)
- val lGI = GIntro_thm
- |> forall_elim (cert f)
- |> fold forall_elim cqs
- |> fold Thm.elim_implies ags
-
- fun mk_call_info (rcfix, rcassm, rcarg) RI =
- let
- val llRI = RI
- |> fold forall_elim cqs
- |> fold (forall_elim o cert o Free) rcfix
- |> fold Thm.elim_implies ags
- |> fold Thm.elim_implies rcassm
-
- val h_assum =
- HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (Logic.all o Free) rcfix
- |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
- |> abstract_over_list (rev qs)
- in
- RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
- end
-
- val RC_infos = map2 mk_call_info RCs RIntro_thms
- in
- ClauseInfo
- {
- no=no,
- cdata=cdata,
- qglr=qglr,
-
- lGI=lGI,
- RCs=RC_infos,
- tree=tree
- }
- end
-
-
-
-
-
-
-
-(* replace this by a table later*)
-fun store_compat_thms 0 thms = []
- | store_compat_thms n thms =
- let
- val (thms1, thms2) = chop n thms
- in
- (thms1 :: store_compat_thms (n - 1) thms2)
- end
-
-(* expects i <= j *)
-fun lookup_compat_thm i j cts =
- nth (nth cts (i - 1)) (j - i)
-
-(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
-(* if j < i, then turn around *)
-fun get_compat_thm thy cts i j ctxi ctxj =
- let
- val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
- val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
-
- val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
- in if j < i then
- let
- val compat = lookup_compat_thm j i cts
- in
- compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
- |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
- |> fold Thm.elim_implies agsj
- |> fold Thm.elim_implies agsi
- |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
- end
- else
- let
- val compat = lookup_compat_thm i j cts
- in
- compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
- |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
- |> fold Thm.elim_implies agsi
- |> fold Thm.elim_implies agsj
- |> Thm.elim_implies (assume lhsi_eq_lhsj)
- |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
- end
- end
-
-
-
-
-(* Generates the replacement lemma in fully quantified form. *)
-fun mk_replacement_lemma thy h ih_elim clause =
- let
- val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
- local open Conv in
- val ih_conv = arg1_conv o arg_conv o arg_conv
- end
-
- val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
-
- val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
- val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
-
- val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
-
- val replace_lemma = (eql RS meta_eq_to_obj_eq)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) h_assums
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
- |> Thm.close_derivation
- in
- replace_lemma
- end
-
-
-fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
- let
- val Globals {h, y, x, fvar, ...} = globals
- val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
- val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
-
- val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
- = mk_clause_context x ctxti cdescj
-
- val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
- val compat = get_compat_thm thy compat_store i j cctxi cctxj
- val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
-
- val RLj_import =
- RLj |> fold forall_elim cqsj'
- |> fold Thm.elim_implies agsj'
- |> fold Thm.elim_implies Ghsj'
-
- val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
- val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
- in
- (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
- |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
- |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
- |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
- |> fold_rev (implies_intr o cprop_of) Ghsj'
- |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
- |> implies_intr (cprop_of y_eq_rhsj'h)
- |> implies_intr (cprop_of lhsi_eq_lhsj')
- |> fold_rev forall_intr (cterm_of thy h :: cqsj')
- end
-
-
-
-fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
- let
- val Globals {x, y, ranT, fvar, ...} = globals
- val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
- val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
- val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
-
- fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
- val existence = fold (curry op COMP o prep_RC) RCs lGI
-
- val P = cterm_of thy (mk_eq (y, rhsC))
- val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
-
- val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
-
- val uniqueness = G_cases
- |> forall_elim (cterm_of thy lhs)
- |> forall_elim (cterm_of thy y)
- |> forall_elim P
- |> Thm.elim_implies G_lhs_y
- |> fold Thm.elim_implies unique_clauses
- |> implies_intr (cprop_of G_lhs_y)
- |> forall_intr (cterm_of thy y)
-
- val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
-
- val exactly_one =
- ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
- |> curry (op COMP) existence
- |> curry (op COMP) uniqueness
- |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
-
- val function_value =
- existence
- |> implies_intr ihyp
- |> implies_intr (cprop_of case_hyp)
- |> forall_intr (cterm_of thy x)
- |> forall_elim (cterm_of thy lhs)
- |> curry (op RS) refl
- in
- (exactly_one, function_value)
- end
-
-
-
-
-fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
- let
- val Globals {h, domT, ranT, x, ...} = globals
- val thy = ProofContext.theory_of ctxt
-
- (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
- HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
- Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
- |> cterm_of thy
-
- val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
- val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
- val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
- |> instantiate' [] [NONE, SOME (cterm_of thy h)]
-
- val _ = Output.debug (K "Proving Replacement lemmas...")
- val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
-
- val _ = Output.debug (K "Proving cases for unique existence...")
- val (ex1s, values) =
- split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
-
- val _ = Output.debug (K "Proving: Graph is a function")
- val graph_is_function = complete
- |> Thm.forall_elim_vars 0
- |> fold (curry op COMP) ex1s
- |> implies_intr (ihyp)
- |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
- |> forall_intr (cterm_of thy x)
- |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
- |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
-
- val goalstate = Conjunction.intr graph_is_function complete
- |> Thm.close_derivation
- |> Goal.protect
- |> fold_rev (implies_intr o cprop_of) compat
- |> implies_intr (cprop_of complete)
- in
- (goalstate, values)
- end
-
-
-fun define_graph Gname fvar domT ranT clauses RCss lthy =
- let
- val GT = domT --> ranT --> boolT
- val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
-
- fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
- let
- fun mk_h_assm (rcfix, rcassm, rcarg) =
- HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (Logic.all o Free) rcfix
- in
- HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
- |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev Logic.all (fvar :: qs)
- end
-
- val G_intros = map2 mk_GIntro clauses RCss
-
- val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
- FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
- in
- ((G, GIntro_thms, G_elim, G_induct), lthy)
- end
-
-
-
-fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
- let
- val f_def =
- Abs ("x", domT, Const ("FunDef.THE_default", ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
- Abs ("y", ranT, G $ Bound 1 $ Bound 0))
- |> Syntax.check_term lthy
-
- val ((f, (_, f_defthm)), lthy) =
- LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
- in
- ((f, f_defthm), lthy)
- end
-
-
-fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
- let
-
- val RT = domT --> domT --> boolT
- val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
-
- fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
- HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev (Logic.all o Free) rcfix
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
-
- val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
-
- val (RIntro_thmss, (R, R_elim, _, lthy)) =
- fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
- in
- ((R, RIntro_thmss, R_elim), lthy)
- end
-
-
-fun fix_globals domT ranT fvar ctxt =
- let
- val ([h, y, x, z, a, D, P, Pbool],ctxt') =
- Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
- in
- (Globals {h = Free (h, domT --> ranT),
- y = Free (y, ranT),
- x = Free (x, domT),
- z = Free (z, domT),
- a = Free (a, domT),
- D = Free (D, domT --> boolT),
- P = Free (P, domT --> boolT),
- Pbool = Free (Pbool, boolT),
- fvar = fvar,
- domT = domT,
- ranT = ranT
- },
- ctxt')
- end
-
-
-
-fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
- let
- fun inst_term t = subst_bound(f, abstract_over (fvar, t))
- in
- (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
- end
-
-
-
-(**********************************************************
- * PROVING THE RULES
- **********************************************************)
-
-fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
- let
- val Globals {domT, z, ...} = globals
-
- fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
- let
- val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
- val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
- in
- ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
- |> (fn it => it COMP graph_is_function)
- |> implies_intr z_smaller
- |> forall_intr (cterm_of thy z)
- |> (fn it => it COMP valthm)
- |> implies_intr lhs_acc
- |> asm_simplify (HOL_basic_ss addsimps [f_iff])
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
- in
- map2 mk_psimp clauses valthms
- end
-
-
-(** Induction rule **)
-
-
-val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
-
-
-fun binder_conv cv ctxt = Conv.arg_conv (Conv.abs_conv (K cv) ctxt);
-
-fun mk_partial_induct_rule thy globals R complete_thm clauses =
- let
- val Globals {domT, x, z, a, P, D, ...} = globals
- val acc_R = mk_acc domT R
-
- val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
- val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
-
- val D_subset = cterm_of thy (Logic.all x
- (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
-
- val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
- Logic.all x
- (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
- HOLogic.mk_Trueprop (D $ z)))))
- |> cterm_of thy
-
-
- (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
- HOLogic.mk_Trueprop (P $ Bound 0)))
- |> cterm_of thy
-
- val aihyp = assume ihyp
-
- fun prove_case clause =
- let
- val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
- qglr = (oqs, _, _, _), ...} = clause
-
- val case_hyp_conv = K (case_hyp RS eq_reflection)
- local open Conv in
- val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
- val sih = fconv_rule (binder_conv (arg1_conv (arg_conv (arg_conv case_hyp_conv))) ctxt) aihyp
- end
-
- fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
- sih |> forall_elim (cterm_of thy rcarg)
- |> Thm.elim_implies llRI
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
- val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
-
- val step = HOLogic.mk_Trueprop (P $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
- |> fold_rev (curry Logic.mk_implies) gs
- |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
-
- val P_lhs = assume step
- |> fold forall_elim cqs
- |> Thm.elim_implies lhs_D
- |> fold Thm.elim_implies ags
- |> fold Thm.elim_implies P_recs
-
- val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
- |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
- |> symmetric (* P lhs == P x *)
- |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
- in
- (res, step)
- end
-
- val (cases, steps) = split_list (map prove_case clauses)
-
- val istep = complete_thm
- |> Thm.forall_elim_vars 0
- |> fold (curry op COMP) cases (* P x *)
- |> implies_intr ihyp
- |> implies_intr (cprop_of x_D)
- |> forall_intr (cterm_of thy x)
-
- val subset_induct_rule =
- acc_subset_induct
- |> (curry op COMP) (assume D_subset)
- |> (curry op COMP) (assume D_dcl)
- |> (curry op COMP) (assume a_D)
- |> (curry op COMP) istep
- |> fold_rev implies_intr steps
- |> implies_intr a_D
- |> implies_intr D_dcl
- |> implies_intr D_subset
-
- val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
-
- val simple_induct_rule =
- subset_induct_rule
- |> forall_intr (cterm_of thy D)
- |> forall_elim (cterm_of thy acc_R)
- |> assume_tac 1 |> Seq.hd
- |> (curry op COMP) (acc_downward
- |> (instantiate' [SOME (ctyp_of thy domT)]
- (map (SOME o cterm_of thy) [R, x, z]))
- |> forall_intr (cterm_of thy z)
- |> forall_intr (cterm_of thy x))
- |> forall_intr (cterm_of thy a)
- |> forall_intr (cterm_of thy P)
- in
- simple_induct_rule
- end
-
-
-
-(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
-fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
- let
- val thy = ProofContext.theory_of ctxt
- val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
- qglr = (oqs, _, _, _), ...} = clause
- val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
- |> fold_rev (curry Logic.mk_implies) gs
- |> cterm_of thy
- in
- Goal.init goal
- |> (SINGLE (resolve_tac [accI] 1)) |> the
- |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
- |> (SINGLE (auto_tac (local_clasimpset_of ctxt))) |> the
- |> Goal.conclude
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
-
-
-
-(** Termination rule **)
-
-val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
-val wf_in_rel = @{thm FunDef.wf_in_rel};
-val in_rel_def = @{thm FunDef.in_rel_def};
-
-fun mk_nest_term_case thy globals R' ihyp clause =
- let
- val Globals {x, z, ...} = globals
- val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
- qglr=(oqs, _, _, _), ...} = clause
-
- val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
-
- fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
- let
- val used = map (fn (ctx,thm) => FundefCtxTree.export_thm thy ctx thm) (u @ sub)
-
- val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
- |> FundefCtxTree.export_term (fixes, assumes)
- |> fold_rev (curry Logic.mk_implies o prop_of) ags
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
-
- val thm = assume hyp
- |> fold forall_elim cqs
- |> fold Thm.elim_implies ags
- |> FundefCtxTree.import_thm thy (fixes, assumes)
- |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
-
- val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
-
- val acc = thm COMP ih_case
- val z_acc_local = acc
- |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
-
- val ethm = z_acc_local
- |> FundefCtxTree.export_thm thy (fixes,
- z_eq_arg :: case_hyp :: ags @ assumes)
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-
- val sub' = sub @ [(([],[]), acc)]
- in
- (sub', (hyp :: hyps, ethm :: thms))
- end
- | step _ _ _ _ = raise Match
- in
- FundefCtxTree.traverse_tree step tree
- end
-
-
-fun mk_nest_term_rule thy globals R R_cases clauses =
- let
- val Globals { domT, x, z, ... } = globals
- val acc_R = mk_acc domT R
-
- val R' = Free ("R", fastype_of R)
-
- val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
- val inrel_R = Const ("FunDef.in_rel", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
-
- val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name "Wellfounded.wfP"}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
-
- (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
- HOLogic.mk_Trueprop (acc_R $ Bound 0)))
- |> cterm_of thy
-
- val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
-
- val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
-
- val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
- in
- R_cases
- |> forall_elim (cterm_of thy z)
- |> forall_elim (cterm_of thy x)
- |> forall_elim (cterm_of thy (acc_R $ z))
- |> curry op COMP (assume R_z_x)
- |> fold_rev (curry op COMP) cases
- |> implies_intr R_z_x
- |> forall_intr (cterm_of thy z)
- |> (fn it => it COMP accI)
- |> implies_intr ihyp
- |> forall_intr (cterm_of thy x)
- |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
- |> curry op RS (assume wfR')
- |> forall_intr_vars
- |> (fn it => it COMP allI)
- |> fold implies_intr hyps
- |> implies_intr wfR'
- |> forall_intr (cterm_of thy R')
- |> forall_elim (cterm_of thy (inrel_R))
- |> curry op RS wf_in_rel
- |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
- |> forall_intr (cterm_of thy Rrel)
- end
-
-
-
-(* Tail recursion (probably very fragile)
- *
- * FIXME:
- * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
- * - Must we really replace the fvar by f here?
- * - Splitting is not configured automatically: Problems with case?
- *)
-fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
- let
- val Globals {domT, ranT, fvar, ...} = globals
-
- val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
-
- val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *)
- Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
- (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
- (fn {prems=[a], ...} =>
- ((rtac (G_induct OF [a]))
- THEN_ALL_NEW (rtac accI)
- THEN_ALL_NEW (etac R_cases)
- THEN_ALL_NEW (asm_full_simp_tac (local_simpset_of octxt))) 1)
-
- val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
-
- fun mk_trsimp clause psimp =
- let
- val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
- val thy = ProofContext.theory_of ctxt
- val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
- val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
- val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
- fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
- in
- Goal.prove ctxt [] [] trsimp
- (fn _ =>
- rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
- THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
- THEN (simp_default_tac (local_simpset_of ctxt) 1)
- THEN (etac not_acc_down 1)
- THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (local_simpset_of ctxt))) 1)
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
- in
- map2 mk_trsimp clauses psimps
- end
-
-
-fun prepare_fundef config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
- let
- val FundefConfig {domintros, tailrec, default=default_str, ...} = config
-
- val fvar = Free (fname, fT)
- val domT = domain_type fT
- val ranT = range_type fT
-
- val default = Syntax.parse_term lthy default_str
- |> TypeInfer.constrain fT |> Syntax.check_term lthy
-
- val (globals, ctxt') = fix_globals domT ranT fvar lthy
-
- val Globals { x, h, ... } = globals
-
- val clauses = map (mk_clause_context x ctxt') abstract_qglrs
-
- val n = length abstract_qglrs
-
- fun build_tree (ClauseContext { ctxt, rhs, ...}) =
- FundefCtxTree.mk_tree (fname, fT) h ctxt rhs
-
- val trees = map build_tree clauses
- val RCss = map find_calls trees
-
- val ((G, GIntro_thms, G_elim, G_induct), lthy) =
- PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
-
- val ((f, f_defthm), lthy) =
- PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
-
- val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
- val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
-
- val ((R, RIntro_thmss, R_elim), lthy) =
- PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
-
- val (_, lthy) =
- LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
-
- val newthy = ProofContext.theory_of lthy
- val clauses = map (transfer_clause_ctx newthy) clauses
-
- val cert = cterm_of (ProofContext.theory_of lthy)
-
- val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
-
- val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
- val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
-
- val compat_store = store_compat_thms n compat
-
- val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
-
- val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
-
- fun mk_partial_rules provedgoal =
- let
- val newthy = theory_of_thm provedgoal (*FIXME*)
-
- val (graph_is_function, complete_thm) =
- provedgoal
- |> Conjunction.elim
- |> apfst (Thm.forall_elim_vars 0)
-
- val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
-
- val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
-
- val simple_pinduct = PROFILE "Proving partial induction rule"
- (mk_partial_induct_rule newthy globals R complete_thm) xclauses
-
-
- val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
-
- val dom_intros = if domintros
- then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
- else NONE
- val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
-
- in
- FundefResult {fs=[f], G=G, R=R, cases=complete_thm,
- psimps=psimps, simple_pinducts=[simple_pinduct],
- termination=total_intro, trsimps=trsimps,
- domintros=dom_intros}
- end
- in
- ((f, goalstate, mk_partial_rules), lthy)
- end
-
-
-end
--- a/src/HOL/Tools/function_package/fundef_datatype.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,330 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_datatype.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-A tactic to prove completeness of datatype patterns.
-*)
-
-signature FUNDEF_DATATYPE =
-sig
- val pat_completeness_tac: Proof.context -> int -> tactic
- val pat_completeness: Proof.context -> Proof.method
- val prove_completeness : theory -> term list -> term -> term list list -> term list list -> thm
-
- val setup : theory -> theory
-
- val add_fun : FundefCommon.fundef_config ->
- (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
- bool -> local_theory -> Proof.context
- val add_fun_cmd : FundefCommon.fundef_config ->
- (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
- bool -> local_theory -> Proof.context
-end
-
-structure FundefDatatype : FUNDEF_DATATYPE =
-struct
-
-open FundefLib
-open FundefCommon
-
-
-fun check_pats ctxt geq =
- let
- fun err str = error (cat_lines ["Malformed definition:",
- str ^ " not allowed in sequential mode.",
- Syntax.string_of_term ctxt geq])
- val thy = ProofContext.theory_of ctxt
-
- fun check_constr_pattern (Bound _) = ()
- | check_constr_pattern t =
- let
- val (hd, args) = strip_comb t
- in
- (((case DatatypePackage.datatype_of_constr thy (fst (dest_Const hd)) of
- SOME _ => ()
- | NONE => err "Non-constructor pattern")
- handle TERM ("dest_Const", _) => err "Non-constructor patterns");
- map check_constr_pattern args;
- ())
- end
-
- val (fname, qs, gs, args, rhs) = split_def ctxt geq
-
- val _ = if not (null gs) then err "Conditional equations" else ()
- val _ = map check_constr_pattern args
-
- (* just count occurrences to check linearity *)
- val _ = if fold (fold_aterms (fn Bound _ => curry (op +) 1 | _ => I)) args 0 > length qs
- then err "Nonlinear patterns" else ()
- in
- ()
- end
-
-
-fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
-fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
-
-fun inst_free var inst thm =
- forall_elim inst (forall_intr var thm)
-
-
-fun inst_case_thm thy x P thm =
- let
- val [Pv, xv] = Term.add_vars (prop_of thm) []
- in
- cterm_instantiate [(cterm_of thy (Var xv), cterm_of thy x),
- (cterm_of thy (Var Pv), cterm_of thy P)] thm
- end
-
-
-fun invent_vars constr i =
- let
- val Ts = binder_types (fastype_of constr)
- val j = i + length Ts
- val is = i upto (j - 1)
- val avs = map2 mk_argvar is Ts
- val pvs = map2 mk_patvar is Ts
- in
- (avs, pvs, j)
- end
-
-
-fun filter_pats thy cons pvars [] = []
- | filter_pats thy cons pvars (([], thm) :: pts) = raise Match
- | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
- case pat of
- Free _ => let val inst = list_comb (cons, pvars)
- in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
- :: (filter_pats thy cons pvars pts) end
- | _ => if fst (strip_comb pat) = cons
- then (pat :: pats, thm) :: (filter_pats thy cons pvars pts)
- else filter_pats thy cons pvars pts
-
-
-fun inst_constrs_of thy (T as Type (name, _)) =
- map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
- (the (DatatypePackage.get_datatype_constrs thy name))
- | inst_constrs_of thy _ = raise Match
-
-
-fun transform_pat thy avars c_assum ([] , thm) = raise Match
- | transform_pat thy avars c_assum (pat :: pats, thm) =
- let
- val (_, subps) = strip_comb pat
- val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
- val a_eqs = map assume eqs
- val c_eq_pat = simplify (HOL_basic_ss addsimps a_eqs) c_assum
- in
- (subps @ pats, fold_rev implies_intr eqs
- (implies_elim thm c_eq_pat))
- end
-
-
-exception COMPLETENESS
-
-fun constr_case thy P idx (v :: vs) pats cons =
- let
- val (avars, pvars, newidx) = invent_vars cons idx
- val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
- val c_assum = assume c_hyp
- val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
- in
- o_alg thy P newidx (avars @ vs) newpats
- |> implies_intr c_hyp
- |> fold_rev (forall_intr o cterm_of thy) avars
- end
- | constr_case _ _ _ _ _ _ = raise Match
-and o_alg thy P idx [] (([], Pthm) :: _) = Pthm
- | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS
- | o_alg thy P idx (v :: vs) pts =
- if forall (is_Free o hd o fst) pts (* Var case *)
- then o_alg thy P idx vs (map (fn (pv :: pats, thm) =>
- (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts)
- else (* Cons case *)
- let
- val T = fastype_of v
- val (tname, _) = dest_Type T
- val {exhaustion=case_thm, ...} = DatatypePackage.the_datatype thy tname
- val constrs = inst_constrs_of thy T
- val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
- in
- inst_case_thm thy v P case_thm
- |> fold (curry op COMP) c_cases
- end
- | o_alg _ _ _ _ _ = raise Match
-
-
-fun prove_completeness thy xs P qss patss =
- let
- fun mk_assum qs pats =
- HOLogic.mk_Trueprop P
- |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)
- |> fold_rev Logic.all qs
- |> cterm_of thy
-
- val hyps = map2 mk_assum qss patss
-
- fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
-
- val assums = map2 inst_hyps hyps qss
- in
- o_alg thy P 2 xs (patss ~~ assums)
- |> fold_rev implies_intr hyps
- end
-
-
-
-fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>
- let
- val thy = ProofContext.theory_of ctxt
- val (vs, subgf) = dest_all_all subgoal
- val (cases, _ $ thesis) = Logic.strip_horn subgf
- handle Bind => raise COMPLETENESS
-
- fun pat_of assum =
- let
- val (qs, imp) = dest_all_all assum
- val prems = Logic.strip_imp_prems imp
- in
- (qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems)
- end
-
- val (qss, x_pats) = split_list (map pat_of cases)
- val xs = map fst (hd x_pats)
- handle Empty => raise COMPLETENESS
-
- val patss = map (map snd) x_pats
-
- val complete_thm = prove_completeness thy xs thesis qss patss
- |> fold_rev (forall_intr o cterm_of thy) vs
- in
- PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st))
- end
- handle COMPLETENESS => no_tac)
-
-
-fun pat_completeness ctxt = SIMPLE_METHOD' (pat_completeness_tac ctxt)
-
-val by_pat_completeness_auto =
- Proof.global_future_terminal_proof
- (Method.Basic (pat_completeness, Position.none),
- SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
-
-fun termination_by method int =
- FundefPackage.termination_proof NONE
- #> Proof.global_future_terminal_proof
- (Method.Basic (method, Position.none), NONE) int
-
-fun mk_catchall fixes arity_of =
- let
- fun mk_eqn ((fname, fT), _) =
- let
- val n = arity_of fname
- val (argTs, rT) = chop n (binder_types fT)
- |> apsnd (fn Ts => Ts ---> body_type fT)
-
- val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
- in
- HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
- Const ("HOL.undefined", rT))
- |> HOLogic.mk_Trueprop
- |> fold_rev Logic.all qs
- end
- in
- map mk_eqn fixes
- end
-
-fun add_catchall ctxt fixes spec =
- let val fqgars = map (split_def ctxt) spec
- val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
- |> AList.lookup (op =) #> the
- in
- spec @ mk_catchall fixes arity_of
- end
-
-fun warn_if_redundant ctxt origs tss =
- let
- fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
-
- val (tss', _) = chop (length origs) tss
- fun check (t, []) = (Output.warning (msg t); [])
- | check (t, s) = s
- in
- (map check (origs ~~ tss'); tss)
- end
-
-
-fun sequential_preproc (config as FundefConfig {sequential, ...}) ctxt fixes spec =
- if sequential then
- let
- val (bnds, eqss) = split_list spec
-
- val eqs = map the_single eqss
-
- val feqs = eqs
- |> tap (check_defs ctxt fixes) (* Standard checks *)
- |> tap (map (check_pats ctxt)) (* More checks for sequential mode *)
-
- val compleqs = add_catchall ctxt fixes feqs (* Completion *)
-
- val spliteqs = warn_if_redundant ctxt feqs
- (FundefSplit.split_all_equations ctxt compleqs)
-
- fun restore_spec thms =
- bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
-
- val spliteqs' = flat (Library.take (length bnds, spliteqs))
- val fnames = map (fst o fst) fixes
- val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
-
- fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
- |> map (map snd)
-
-
- val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
-
- (* using theorem names for case name currently disabled *)
- val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es))
- (bnds' ~~ spliteqs)
- |> flat
- in
- (flat spliteqs, restore_spec, sort, case_names)
- end
- else
- FundefCommon.empty_preproc check_defs config ctxt fixes spec
-
-val setup =
- Method.setup @{binding pat_completeness} (Scan.succeed pat_completeness)
- "Completeness prover for datatype patterns"
- #> Context.theory_map (FundefCommon.set_preproc sequential_preproc)
-
-
-val fun_config = FundefConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*),
- domintros=false, tailrec=false }
-
-fun gen_fun add config fixes statements int lthy =
- let val group = serial_string () in
- lthy
- |> LocalTheory.set_group group
- |> add fixes statements config
- |> by_pat_completeness_auto int
- |> LocalTheory.restore
- |> LocalTheory.set_group group
- |> termination_by (FundefCommon.get_termination_prover lthy) int
- end;
-
-val add_fun = gen_fun FundefPackage.add_fundef
-val add_fun_cmd = gen_fun FundefPackage.add_fundef_cmd
-
-
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
- OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
- (fundef_parser fun_config
- >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
-
-end
-
-end
--- a/src/HOL/Tools/function_package/fundef_lib.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,176 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_lib.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Some fairly general functions that should probably go somewhere else...
-*)
-
-structure FundefLib = struct
-
-fun map_option f NONE = NONE
- | map_option f (SOME x) = SOME (f x);
-
-fun fold_option f NONE y = y
- | fold_option f (SOME x) y = f x y;
-
-fun fold_map_option f NONE y = (NONE, y)
- | fold_map_option f (SOME x) y = apfst SOME (f x y);
-
-(* Ex: "The variable" ^ plural " is" "s are" vs *)
-fun plural sg pl [x] = sg
- | plural sg pl _ = pl
-
-(* lambda-abstracts over an arbitrarily nested tuple
- ==> hologic.ML? *)
-fun tupled_lambda vars t =
- case vars of
- (Free v) => lambda (Free v) t
- | (Var v) => lambda (Var v) t
- | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
- (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
- | _ => raise Match
-
-
-fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
- let
- val (n, body) = Term.dest_abs a
- in
- (Free (n, T), body)
- end
- | dest_all _ = raise Match
-
-
-(* Removes all quantifiers from a term, replacing bound variables by frees. *)
-fun dest_all_all (t as (Const ("all",_) $ _)) =
- let
- val (v,b) = dest_all t
- val (vs, b') = dest_all_all b
- in
- (v :: vs, b')
- end
- | dest_all_all t = ([],t)
-
-
-(* FIXME: similar to Variable.focus *)
-fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
- let
- val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
- val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
-
- val (n'', body) = Term.dest_abs (n', T, b)
- val _ = (n' = n'') orelse error "dest_all_ctx"
- (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
-
- val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
- in
- (ctx'', (n', T) :: vs, bd)
- end
- | dest_all_all_ctx ctx t =
- (ctx, [], t)
-
-
-fun map3 _ [] [] [] = []
- | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
- | map3 _ _ _ _ = raise Library.UnequalLengths;
-
-fun map4 _ [] [] [] [] = []
- | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
- | map4 _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map6 _ [] [] [] [] [] [] = []
- | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
- | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map7 _ [] [] [] [] [] [] [] = []
- | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
- | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-
-
-(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
-(* ==> library *)
-fun unordered_pairs [] = []
- | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
-
-
-(* Replaces Frees by name. Works with loose Bounds. *)
-fun replace_frees assoc =
- map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
- | t => t)
-
-
-fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
- | rename_bound n _ = raise Match
-
-fun mk_forall_rename (n, v) =
- rename_bound n o Logic.all v
-
-fun forall_intr_rename (n, cv) thm =
- let
- val allthm = forall_intr cv thm
- val (_ $ abs) = prop_of allthm
- in
- Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
- end
-
-
-(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
-fun frees_in_term ctxt t =
- Term.add_frees t []
- |> filter_out (Variable.is_fixed ctxt o fst)
- |> rev
-
-
-datatype proof_attempt = Solved of thm | Stuck of thm | Fail
-
-fun try_proof cgoal tac =
- case SINGLE tac (Goal.init cgoal) of
- NONE => Fail
- | SOME st => if Thm.no_prems st then Solved (Goal.finish st) else Stuck st
-
-
-fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) =
- if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
- | dest_binop_list _ t = [ t ]
-
-
-(* separate two parts in a +-expression:
- "a + b + c + d + e" --> "(a + b + d) + (c + e)"
-
- Here, + can be any binary operation that is AC.
-
- cn - The name of the binop-constructor (e.g. @{const_name Un})
- ac - the AC rewrite rules for cn
- is - the list of indices of the expressions that should become the first part
- (e.g. [0,1,3] in the above example)
-*)
-
-fun regroup_conv neu cn ac is ct =
- let
- val mk = HOLogic.mk_binop cn
- val t = term_of ct
- val xs = dest_binop_list cn t
- val js = 0 upto (length xs) - 1 \\ is
- val ty = fastype_of t
- val thy = theory_of_cterm ct
- in
- Goal.prove_internal []
- (cterm_of thy
- (Logic.mk_equals (t,
- if is = []
- then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
- else if js = []
- then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
- else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
- (K (rewrite_goals_tac ac
- THEN rtac Drule.reflexive_thm 1))
- end
-
-(* instance for unions *)
-fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Un}
- (map (fn t => t RS eq_reflection) (@{thms "Un_ac"} @
- @{thms "Un_empty_right"} @
- @{thms "Un_empty_left"})) t
-
-
-end
--- a/src/HOL/Tools/function_package/fundef_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,226 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_package.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Isar commands.
-*)
-
-signature FUNDEF_PACKAGE =
-sig
- val add_fundef : (binding * typ option * mixfix) list
- -> (Attrib.binding * term) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
- val add_fundef_cmd : (binding * string option * mixfix) list
- -> (Attrib.binding * string) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
-
- val termination_proof : term option -> local_theory -> Proof.state
- val termination_proof_cmd : string option -> local_theory -> Proof.state
- val termination : term option -> local_theory -> Proof.state
- val termination_cmd : string option -> local_theory -> Proof.state
-
- val setup : theory -> theory
- val get_congs : Proof.context -> thm list
-end
-
-
-structure FundefPackage : FUNDEF_PACKAGE =
-struct
-
-open FundefLib
-open FundefCommon
-
-val simp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Code.add_default_eqn_attribute,
- Nitpick_Const_Simp_Thms.add,
- Quickcheck_RecFun_Simp_Thms.add]
-
-val psimp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Nitpick_Const_Psimp_Thms.add]
-
-fun note_theorem ((name, atts), ths) =
- LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
-
-fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
-
-fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
- let
- val spec = post simps
- |> map (apfst (apsnd (fn ats => moreatts @ ats)))
- |> map (apfst (apfst extra_qualify))
-
- val (saved_spec_simps, lthy) =
- fold_map (LocalTheory.note Thm.generatedK) spec lthy
-
- val saved_simps = flat (map snd saved_spec_simps)
- val simps_by_f = sort saved_simps
-
- fun add_for_f fname simps =
- note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
- in
- (saved_simps,
- fold2 add_for_f fnames simps_by_f lthy)
- end
-
-fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
- let
- val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
- val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
- val fixes = map (apfst (apfst Binding.name_of)) fixes0;
- val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
- val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
-
- val defname = mk_defname fixes
-
- val ((goalstate, cont), lthy) =
- FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
-
- fun afterqed [[proof]] lthy =
- let
- val FundefResult {fs, R, psimps, trsimps, simple_pinducts, termination,
- domintros, cases, ...} =
- cont (Thm.close_derivation proof)
-
- val fnames = map (fst o fst) fixes
- val qualify = Long_Name.qualify defname
- val addsmps = add_simps fnames post sort_cont
-
- val (((psimps', pinducts'), (_, [termination'])), lthy) =
- lthy
- |> addsmps (Binding.qualify false "partial") "psimps"
- psimp_attribs psimps
- ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
- ||>> note_theorem ((qualify "pinduct",
- [Attrib.internal (K (RuleCases.case_names cnames)),
- Attrib.internal (K (RuleCases.consumes 1)),
- Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
- ||>> note_theorem ((qualify "termination", []), [termination])
- ||> (snd o note_theorem ((qualify "cases",
- [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
- ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
-
- val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
- pinducts=snd pinducts', termination=termination',
- fs=fs, R=R, defname=defname }
- val _ =
- if not is_external then ()
- else Specification.print_consts lthy (K false) (map fst fixes)
- in
- lthy
- |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
- end
- in
- lthy
- |> is_external ? LocalTheory.set_group (serial_string ())
- |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
- |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
- end
-
-val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
-val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
-
-fun gen_termination_proof prep_term raw_term_opt lthy =
- let
- val term_opt = Option.map (prep_term lthy) raw_term_opt
- val data = the (case term_opt of
- SOME t => (import_fundef_data t lthy
- handle Option.Option =>
- error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
- | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
-
- val FundefCtxData { termination, R, add_simps, case_names, psimps,
- pinducts, defname, ...} = data
- val domT = domain_type (fastype_of R)
- val goal = HOLogic.mk_Trueprop
- (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
- fun afterqed [[totality]] lthy =
- let
- val totality = Thm.close_derivation totality
- val remove_domain_condition =
- full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
- val tsimps = map remove_domain_condition psimps
- val tinduct = map remove_domain_condition pinducts
- val qualify = Long_Name.qualify defname;
- in
- lthy
- |> add_simps I "simps" simp_attribs tsimps |> snd
- |> note_theorem
- ((qualify "induct",
- [Attrib.internal (K (RuleCases.case_names case_names))]),
- tinduct) |> snd
- end
- in
- lthy
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
- [([Goal.norm_result termination], [])])] |> snd
- |> Proof.theorem_i NONE afterqed [[(goal, [])]]
- end
-
-val termination_proof = gen_termination_proof Syntax.check_term;
-val termination_proof_cmd = gen_termination_proof Syntax.read_term;
-
-fun termination term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof term_opt;
-
-fun termination_cmd term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof_cmd term_opt;
-
-
-(* Datatype hook to declare datatype congs as "fundef_congs" *)
-
-
-fun add_case_cong n thy =
- Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
- (DatatypePackage.get_datatype thy n |> the
- |> #case_cong
- |> safe_mk_meta_eq)))
- thy
-
-val setup_case_cong = DatatypePackage.interpretation (K (fold add_case_cong))
-
-
-(* setup *)
-
-val setup =
- Attrib.setup @{binding fundef_cong}
- (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
- "declaration of congruence rule for function definitions"
- #> setup_case_cong
- #> FundefRelation.setup
- #> FundefCommon.TerminationSimps.setup
-
-val get_congs = FundefCtxTree.get_fundef_congs
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
- OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
- (fundef_parser default_config
- >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
-
-val _ =
- OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
- (Scan.option P.term >> termination_cmd);
-
-end;
-
-
-end
--- a/src/HOL/Tools/function_package/induction_scheme.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,405 +0,0 @@
-(* Title: HOL/Tools/function_package/induction_scheme.ML
- Author: Alexander Krauss, TU Muenchen
-
-A method to prove induction schemes.
-*)
-
-signature INDUCTION_SCHEME =
-sig
- val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic)
- -> Proof.context -> thm list -> tactic
- val induct_scheme_tac : Proof.context -> thm list -> tactic
- val setup : theory -> theory
-end
-
-
-structure InductionScheme : INDUCTION_SCHEME =
-struct
-
-open FundefLib
-
-
-type rec_call_info = int * (string * typ) list * term list * term list
-
-datatype scheme_case =
- SchemeCase of
- {
- bidx : int,
- qs: (string * typ) list,
- oqnames: string list,
- gs: term list,
- lhs: term list,
- rs: rec_call_info list
- }
-
-datatype scheme_branch =
- SchemeBranch of
- {
- P : term,
- xs: (string * typ) list,
- ws: (string * typ) list,
- Cs: term list
- }
-
-datatype ind_scheme =
- IndScheme of
- {
- T: typ, (* sum of products *)
- branches: scheme_branch list,
- cases: scheme_case list
- }
-
-val ind_atomize = MetaSimplifier.rewrite true @{thms induct_atomize}
-val ind_rulify = MetaSimplifier.rewrite true @{thms induct_rulify}
-
-fun meta thm = thm RS eq_reflection
-
-val sum_prod_conv = MetaSimplifier.rewrite true
- (map meta (@{thm split_conv} :: @{thms sum.cases}))
-
-fun term_conv thy cv t =
- cv (cterm_of thy t)
- |> prop_of |> Logic.dest_equals |> snd
-
-fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T))
-
-fun dest_hhf ctxt t =
- let
- val (ctxt', vars, imp) = dest_all_all_ctx ctxt t
- in
- (ctxt', vars, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp)
- end
-
-
-fun mk_scheme' ctxt cases concl =
- let
- fun mk_branch concl =
- let
- val (ctxt', ws, Cs, _ $ Pxs) = dest_hhf ctxt concl
- val (P, xs) = strip_comb Pxs
- in
- SchemeBranch { P=P, xs=map dest_Free xs, ws=ws, Cs=Cs }
- end
-
- val (branches, cases') = (* correction *)
- case Logic.dest_conjunction_list concl of
- [conc] =>
- let
- val _ $ Pxs = Logic.strip_assums_concl conc
- val (P, _) = strip_comb Pxs
- val (cases', conds) = take_prefix (Term.exists_subterm (curry op aconv P)) cases
- val concl' = fold_rev (curry Logic.mk_implies) conds conc
- in
- ([mk_branch concl'], cases')
- end
- | concls => (map mk_branch concls, cases)
-
- fun mk_case premise =
- let
- val (ctxt', qs, prems, _ $ Plhs) = dest_hhf ctxt premise
- val (P, lhs) = strip_comb Plhs
-
- fun bidx Q = find_index (fn SchemeBranch {P=P',...} => Q aconv P') branches
-
- fun mk_rcinfo pr =
- let
- val (ctxt'', Gvs, Gas, _ $ Phyp) = dest_hhf ctxt' pr
- val (P', rcs) = strip_comb Phyp
- in
- (bidx P', Gvs, Gas, rcs)
- end
-
- fun is_pred v = exists (fn SchemeBranch {P,...} => v aconv P) branches
-
- val (gs, rcprs) =
- take_prefix (not o Term.exists_subterm is_pred) prems
- in
- SchemeCase {bidx=bidx P, qs=qs, oqnames=map fst qs(*FIXME*), gs=gs, lhs=lhs, rs=map mk_rcinfo rcprs}
- end
-
- fun PT_of (SchemeBranch { xs, ...}) =
- foldr1 HOLogic.mk_prodT (map snd xs)
-
- val ST = BalancedTree.make (uncurry SumTree.mk_sumT) (map PT_of branches)
- in
- IndScheme {T=ST, cases=map mk_case cases', branches=branches }
- end
-
-
-
-fun mk_completeness ctxt (IndScheme {cases, branches, ...}) bidx =
- let
- val SchemeBranch { xs, ws, Cs, ... } = nth branches bidx
- val relevant_cases = filter (fn SchemeCase {bidx=bidx', ...} => bidx' = bidx) cases
-
- val allqnames = fold (fn SchemeCase {qs, ...} => fold (insert (op =) o Free) qs) relevant_cases []
- val (Pbool :: xs') = map Free (Variable.variant_frees ctxt allqnames (("P", HOLogic.boolT) :: xs))
- val Cs' = map (Pattern.rewrite_term (ProofContext.theory_of ctxt) (filter_out (op aconv) (map Free xs ~~ xs')) []) Cs
-
- fun mk_case (SchemeCase {qs, oqnames, gs, lhs, ...}) =
- HOLogic.mk_Trueprop Pbool
- |> fold_rev (fn x_l => curry Logic.mk_implies (HOLogic.mk_Trueprop(HOLogic.mk_eq x_l)))
- (xs' ~~ lhs)
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
- in
- HOLogic.mk_Trueprop Pbool
- |> fold_rev (curry Logic.mk_implies o mk_case) relevant_cases
- |> fold_rev (curry Logic.mk_implies) Cs'
- |> fold_rev (Logic.all o Free) ws
- |> fold_rev mk_forall_rename (map fst xs ~~ xs')
- |> mk_forall_rename ("P", Pbool)
- end
-
-fun mk_wf ctxt R (IndScheme {T, ...}) =
- HOLogic.Trueprop $ (Const (@{const_name "wf"}, mk_relT T --> HOLogic.boolT) $ R)
-
-fun mk_ineqs R (IndScheme {T, cases, branches}) =
- let
- fun inject i ts =
- SumTree.mk_inj T (length branches) (i + 1) (foldr1 HOLogic.mk_prod ts)
-
- val thesis = Free ("thesis", HOLogic.boolT) (* FIXME *)
-
- fun mk_pres bdx args =
- let
- val SchemeBranch { xs, ws, Cs, ... } = nth branches bdx
- fun replace (x, v) t = betapply (lambda (Free x) t, v)
- val Cs' = map (fold replace (xs ~~ args)) Cs
- val cse =
- HOLogic.mk_Trueprop thesis
- |> fold_rev (curry Logic.mk_implies) Cs'
- |> fold_rev (Logic.all o Free) ws
- in
- Logic.mk_implies (cse, HOLogic.mk_Trueprop thesis)
- end
-
- fun f (SchemeCase {bidx, qs, oqnames, gs, lhs, rs, ...}) =
- let
- fun g (bidx', Gvs, Gas, rcarg) =
- let val export =
- fold_rev (curry Logic.mk_implies) Gas
- #> fold_rev (curry Logic.mk_implies) gs
- #> fold_rev (Logic.all o Free) Gvs
- #> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
- in
- (HOLogic.mk_mem (HOLogic.mk_prod (inject bidx' rcarg, inject bidx lhs), R)
- |> HOLogic.mk_Trueprop
- |> export,
- mk_pres bidx' rcarg
- |> export
- |> Logic.all thesis)
- end
- in
- map g rs
- end
- in
- map f cases
- end
-
-
-fun mk_hol_imp a b = HOLogic.imp $ a $ b
-
-fun mk_ind_goal thy branches =
- let
- fun brnch (SchemeBranch { P, xs, ws, Cs, ... }) =
- HOLogic.mk_Trueprop (list_comb (P, map Free xs))
- |> fold_rev (curry Logic.mk_implies) Cs
- |> fold_rev (Logic.all o Free) ws
- |> term_conv thy ind_atomize
- |> ObjectLogic.drop_judgment thy
- |> tupled_lambda (foldr1 HOLogic.mk_prod (map Free xs))
- in
- SumTree.mk_sumcases HOLogic.boolT (map brnch branches)
- end
-
-
-fun mk_induct_rule ctxt R x complete_thms wf_thm ineqss (IndScheme {T, cases=scases, branches}) =
- let
- val n = length branches
-
- val scases_idx = map_index I scases
-
- fun inject i ts =
- SumTree.mk_inj T n (i + 1) (foldr1 HOLogic.mk_prod ts)
- val P_of = nth (map (fn (SchemeBranch { P, ... }) => P) branches)
-
- val thy = ProofContext.theory_of ctxt
- val cert = cterm_of thy
-
- val P_comp = mk_ind_goal thy branches
-
- (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
- val ihyp = Term.all T $ Abs ("z", T,
- Logic.mk_implies
- (HOLogic.mk_Trueprop (
- Const ("op :", HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT)
- $ (HOLogic.pair_const T T $ Bound 0 $ x)
- $ R),
- HOLogic.mk_Trueprop (P_comp $ Bound 0)))
- |> cert
-
- val aihyp = assume ihyp
-
- (* Rule for case splitting along the sum types *)
- val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
- val pats = map_index (uncurry inject) xss
- val sum_split_rule = FundefDatatype.prove_completeness thy [x] (P_comp $ x) xss (map single pats)
-
- fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
- let
- val fxs = map Free xs
- val branch_hyp = assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat))))
-
- val C_hyps = map (cert #> assume) Cs
-
- val (relevant_cases, ineqss') = filter (fn ((_, SchemeCase {bidx=bidx', ...}), _) => bidx' = bidx) (scases_idx ~~ ineqss)
- |> split_list
-
- fun prove_case (cidx, SchemeCase {qs, oqnames, gs, lhs, rs, ...}) ineq_press =
- let
- val case_hyps = map (assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs)
-
- val cqs = map (cert o Free) qs
- val ags = map (assume o cert) gs
-
- val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps)
- val sih = full_simplify replace_x_ss aihyp
-
- fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) =
- let
- val cGas = map (assume o cert) Gas
- val cGvs = map (cert o Free) Gvs
- val import = fold forall_elim (cqs @ cGvs)
- #> fold Thm.elim_implies (ags @ cGas)
- val ipres = pres
- |> forall_elim (cert (list_comb (P_of idx, rcargs)))
- |> import
- in
- sih |> forall_elim (cert (inject idx rcargs))
- |> Thm.elim_implies (import ineq) (* Psum rcargs *)
- |> Conv.fconv_rule sum_prod_conv
- |> Conv.fconv_rule ind_rulify
- |> (fn th => th COMP ipres) (* P rs *)
- |> fold_rev (implies_intr o cprop_of) cGas
- |> fold_rev forall_intr cGvs
- end
-
- val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *)
-
- val step = HOLogic.mk_Trueprop (list_comb (P, lhs))
- |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev (Logic.all o Free) qs
- |> cert
-
- val Plhs_to_Pxs_conv =
- foldl1 (uncurry Conv.combination_conv)
- (Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps)
-
- val res = assume step
- |> fold forall_elim cqs
- |> fold Thm.elim_implies ags
- |> fold Thm.elim_implies P_recs (* P lhs *)
- |> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *)
- |> fold_rev (implies_intr o cprop_of) (ags @ case_hyps)
- |> fold_rev forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
- in
- (res, (cidx, step))
- end
-
- val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss')
-
- val bstep = complete_thm
- |> forall_elim (cert (list_comb (P, fxs)))
- |> fold (forall_elim o cert) (fxs @ map Free ws)
- |> fold Thm.elim_implies C_hyps (* FIXME: optimization using rotate_prems *)
- |> fold Thm.elim_implies cases (* P xs *)
- |> fold_rev (implies_intr o cprop_of) C_hyps
- |> fold_rev (forall_intr o cert o Free) ws
-
- val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x))
- |> Goal.init
- |> (MetaSimplifier.rewrite_goals_tac (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.cases}))
- THEN CONVERSION ind_rulify 1)
- |> Seq.hd
- |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
- |> Goal.finish
- |> implies_intr (cprop_of branch_hyp)
- |> fold_rev (forall_intr o cert) fxs
- in
- (Pxs, steps)
- end
-
- val (branches, steps) = split_list (map_index prove_branch (branches ~~ (complete_thms ~~ pats)))
- |> apsnd flat
-
- val istep = sum_split_rule
- |> fold (fn b => fn th => Drule.compose_single (b, 1, th)) branches
- |> implies_intr ihyp
- |> forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *)
-
- val induct_rule =
- @{thm "wf_induct_rule"}
- |> (curry op COMP) wf_thm
- |> (curry op COMP) istep
-
- val steps_sorted = map snd (sort (int_ord o pairself fst) steps)
- in
- (steps_sorted, induct_rule)
- end
-
-
-fun mk_ind_tac comp_tac pres_tac term_tac ctxt facts = (ALLGOALS (Method.insert_tac facts)) THEN HEADGOAL
-(SUBGOAL (fn (t, i) =>
- let
- val (ctxt', _, cases, concl) = dest_hhf ctxt t
- val scheme as IndScheme {T=ST, branches, ...} = mk_scheme' ctxt' cases concl
-(* val _ = Output.tracing (makestring scheme)*)
- val ([Rn,xn], ctxt'') = Variable.variant_fixes ["R","x"] ctxt'
- val R = Free (Rn, mk_relT ST)
- val x = Free (xn, ST)
- val cert = cterm_of (ProofContext.theory_of ctxt)
-
- val ineqss = mk_ineqs R scheme
- |> map (map (pairself (assume o cert)))
- val complete = map (mk_completeness ctxt scheme #> cert #> assume) (0 upto (length branches - 1))
- val wf_thm = mk_wf ctxt R scheme |> cert |> assume
-
- val (descent, pres) = split_list (flat ineqss)
- val newgoals = complete @ pres @ wf_thm :: descent
-
- val (steps, indthm) = mk_induct_rule ctxt'' R x complete wf_thm ineqss scheme
-
- fun project (i, SchemeBranch {xs, ...}) =
- let
- val inst = cert (SumTree.mk_inj ST (length branches) (i + 1) (foldr1 HOLogic.mk_prod (map Free xs)))
- in
- indthm |> Drule.instantiate' [] [SOME inst]
- |> simplify SumTree.sumcase_split_ss
- |> Conv.fconv_rule ind_rulify
-(* |> (fn thm => (Output.tracing (makestring thm); thm))*)
- end
-
- val res = Conjunction.intr_balanced (map_index project branches)
- |> fold_rev implies_intr (map cprop_of newgoals @ steps)
- |> (fn thm => Thm.generalize ([], [Rn]) (Thm.maxidx_of thm + 1) thm)
-
- val nbranches = length branches
- val npres = length pres
- in
- Thm.compose_no_flatten false (res, length newgoals) i
- THEN term_tac (i + nbranches + npres)
- THEN (EVERY (map (TRY o pres_tac) ((i + nbranches + npres - 1) downto (i + nbranches))))
- THEN (EVERY (map (TRY o comp_tac) ((i + nbranches - 1) downto i)))
- end))
-
-
-fun induct_scheme_tac ctxt =
- mk_ind_tac (K all_tac) (assume_tac APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt;
-
-val setup =
- Method.setup @{binding induct_scheme} (Scan.succeed (RAW_METHOD o induct_scheme_tac))
- "proves an induction principle"
-
-end
--- a/src/HOL/Tools/function_package/inductive_wrap.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,68 +0,0 @@
-(* Title: HOL/Tools/function_package/inductive_wrap.ML
- ID: $Id$
- Author: Alexander Krauss, TU Muenchen
-
-
-A wrapper around the inductive package, restoring the quantifiers in
-the introduction and elimination rules.
-*)
-
-signature FUNDEF_INDUCTIVE_WRAP =
-sig
- val inductive_def : term list
- -> ((bstring * typ) * mixfix) * local_theory
- -> thm list * (term * thm * thm * local_theory)
-end
-
-structure FundefInductiveWrap: FUNDEF_INDUCTIVE_WRAP =
-struct
-
-open FundefLib
-
-fun requantify ctxt lfix orig_def thm =
- let
- val (qs, t) = dest_all_all orig_def
- val thy = theory_of_thm thm
- val frees = frees_in_term ctxt t
- |> remove (op =) lfix
- val vars = Term.add_vars (prop_of thm) [] |> rev
-
- val varmap = frees ~~ vars
- in
- fold_rev (fn Free (n, T) =>
- forall_intr_rename (n, cterm_of thy (Var (the_default (("",0), T) (AList.lookup (op =) varmap (n, T))))))
- qs
- thm
- end
-
-
-
-fun inductive_def defs (((R, T), mixfix), lthy) =
- let
- val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) =
- InductivePackage.add_inductive_i
- {quiet_mode = false,
- verbose = ! Toplevel.debug,
- kind = Thm.internalK,
- alt_name = Binding.empty,
- coind = false,
- no_elim = false,
- no_ind = false,
- skip_mono = true,
- fork_mono = false}
- [((Binding.name R, T), NoSyn)] (* the relation *)
- [] (* no parameters *)
- (map (fn t => (Attrib.empty_binding, t)) defs) (* the intros *)
- [] (* no special monos *)
- lthy
-
- val intrs = map2 (requantify lthy (R, T)) defs intrs_gen
-
- val elim = elim_gen
- |> forall_intr_vars (* FIXME... *)
-
- in
- (intrs, (Rdef, elim, induct, lthy))
- end
-
-end
--- a/src/HOL/Tools/function_package/lexicographic_order.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,230 +0,0 @@
-(* Title: HOL/Tools/function_package/lexicographic_order.ML
- Author: Lukas Bulwahn, TU Muenchen
-
-Method for termination proofs with lexicographic orderings.
-*)
-
-signature LEXICOGRAPHIC_ORDER =
-sig
- val lex_order_tac : Proof.context -> tactic -> tactic
- val lexicographic_order_tac : Proof.context -> tactic
- val lexicographic_order : Proof.context -> Proof.method
-
- val setup: theory -> theory
-end
-
-structure LexicographicOrder : LEXICOGRAPHIC_ORDER =
-struct
-
-open FundefLib
-
-(** General stuff **)
-
-fun mk_measures domT mfuns =
- let
- val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))
- val mlexT = (domT --> HOLogic.natT) --> relT --> relT
- fun mk_ms [] = Const (@{const_name Set.empty}, relT)
- | mk_ms (f::fs) =
- Const (@{const_name "mlex_prod"}, mlexT) $ f $ mk_ms fs
- in
- mk_ms mfuns
- end
-
-fun del_index n [] = []
- | del_index n (x :: xs) =
- if n > 0 then x :: del_index (n - 1) xs else xs
-
-fun transpose ([]::_) = []
- | transpose xss = map hd xss :: transpose (map tl xss)
-
-(** Matrix cell datatype **)
-
-datatype cell = Less of thm| LessEq of (thm * thm) | None of (thm * thm) | False of thm;
-
-fun is_Less (Less _) = true
- | is_Less _ = false
-
-fun is_LessEq (LessEq _) = true
- | is_LessEq _ = false
-
-fun pr_cell (Less _ ) = " < "
- | pr_cell (LessEq _) = " <="
- | pr_cell (None _) = " ? "
- | pr_cell (False _) = " F "
-
-
-(** Proof attempts to build the matrix **)
-
-fun dest_term (t : term) =
- let
- val (vars, prop) = FundefLib.dest_all_all t
- val (prems, concl) = Logic.strip_horn prop
- val (lhs, rhs) = concl
- |> HOLogic.dest_Trueprop
- |> HOLogic.dest_mem |> fst
- |> HOLogic.dest_prod
- in
- (vars, prems, lhs, rhs)
- end
-
-fun mk_goal (vars, prems, lhs, rhs) rel =
- let
- val concl = HOLogic.mk_binrel rel (lhs, rhs) |> HOLogic.mk_Trueprop
- in
- fold_rev Logic.all vars (Logic.list_implies (prems, concl))
- end
-
-fun prove thy solve_tac t =
- cterm_of thy t |> Goal.init
- |> SINGLE solve_tac |> the
-
-fun mk_cell (thy : theory) solve_tac (vars, prems, lhs, rhs) mfun =
- let
- val goals = cterm_of thy o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs)
- in
- case try_proof (goals @{const_name HOL.less}) solve_tac of
- Solved thm => Less thm
- | Stuck thm =>
- (case try_proof (goals @{const_name HOL.less_eq}) solve_tac of
- Solved thm2 => LessEq (thm2, thm)
- | Stuck thm2 =>
- if prems_of thm2 = [HOLogic.Trueprop $ HOLogic.false_const] then False thm2
- else None (thm2, thm)
- | _ => raise Match) (* FIXME *)
- | _ => raise Match
- end
-
-
-(** Search algorithms **)
-
-fun check_col ls = forall (fn c => is_Less c orelse is_LessEq c) ls andalso not (forall (is_LessEq) ls)
-
-fun transform_table table col = table |> filter_out (fn x => is_Less (nth x col)) |> map (del_index col)
-
-fun transform_order col order = map (fn x => if x >= col then x + 1 else x) order
-
-(* simple depth-first search algorithm for the table *)
-fun search_table table =
- case table of
- [] => SOME []
- | _ =>
- let
- val col = find_index (check_col) (transpose table)
- in case col of
- ~1 => NONE
- | _ =>
- let
- val order_opt = (table, col) |-> transform_table |> search_table
- in case order_opt of
- NONE => NONE
- | SOME order =>SOME (col :: transform_order col order)
- end
- end
-
-(** Proof Reconstruction **)
-
-(* prove row :: cell list -> tactic *)
-fun prove_row (Less less_thm :: _) =
- (rtac @{thm "mlex_less"} 1)
- THEN PRIMITIVE (Thm.elim_implies less_thm)
- | prove_row (LessEq (lesseq_thm, _) :: tail) =
- (rtac @{thm "mlex_leq"} 1)
- THEN PRIMITIVE (Thm.elim_implies lesseq_thm)
- THEN prove_row tail
- | prove_row _ = sys_error "lexicographic_order"
-
-
-(** Error reporting **)
-
-fun pr_table table = writeln (cat_lines (map (fn r => concat (map pr_cell r)) table))
-
-fun pr_goals ctxt st =
- Display.pretty_goals_aux (Syntax.pp ctxt) Markup.none (true, false) (Thm.nprems_of st) st
- |> Pretty.chunks
- |> Pretty.string_of
-
-fun row_index i = chr (i + 97)
-fun col_index j = string_of_int (j + 1)
-
-fun pr_unprovable_cell _ ((i,j), Less _) = ""
- | pr_unprovable_cell ctxt ((i,j), LessEq (_, st)) =
- "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
- | pr_unprovable_cell ctxt ((i,j), None (st_leq, st_less)) =
- "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st_less
- ^ "\n(" ^ row_index i ^ ", " ^ col_index j ^ ", <=):\n" ^ pr_goals ctxt st_leq
- | pr_unprovable_cell ctxt ((i,j), False st) =
- "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
-
-fun pr_unprovable_subgoals ctxt table =
- table
- |> map_index (fn (i,cs) => map_index (fn (j,x) => ((i,j), x)) cs)
- |> flat
- |> map (pr_unprovable_cell ctxt)
-
-fun no_order_msg ctxt table tl measure_funs =
- let
- val prterm = Syntax.string_of_term ctxt
- fun pr_fun t i = string_of_int i ^ ") " ^ prterm t
-
- fun pr_goal t i =
- let
- val (_, _, lhs, rhs) = dest_term t
- in (* also show prems? *)
- i ^ ") " ^ prterm rhs ^ " ~> " ^ prterm lhs
- end
-
- val gc = map (fn i => chr (i + 96)) (1 upto length table)
- val mc = 1 upto length measure_funs
- val tstr = "Result matrix:" :: (" " ^ concat (map (enclose " " " " o string_of_int) mc))
- :: map2 (fn r => fn i => i ^ ": " ^ concat (map pr_cell r)) table gc
- val gstr = "Calls:" :: map2 (prefix " " oo pr_goal) tl gc
- val mstr = "Measures:" :: map2 (prefix " " oo pr_fun) measure_funs mc
- val ustr = "Unfinished subgoals:" :: pr_unprovable_subgoals ctxt table
- in
- cat_lines (ustr @ gstr @ mstr @ tstr @ ["", "Could not find lexicographic termination order."])
- end
-
-(** The Main Function **)
-
-fun lex_order_tac ctxt solve_tac (st: thm) =
- let
- val thy = ProofContext.theory_of ctxt
- val ((trueprop $ (wf $ rel)) :: tl) = prems_of st
-
- val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel))
-
- val measure_funs = MeasureFunctions.get_measure_functions ctxt domT (* 1: generate measures *)
-
- (* 2: create table *)
- val table = map (fn t => map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl
-
- val order = the (search_table table) (* 3: search table *)
- handle Option => error (no_order_msg ctxt table tl measure_funs)
-
- val clean_table = map (fn x => map (nth x) order) table
-
- val relation = mk_measures domT (map (nth measure_funs) order)
- val _ = writeln ("Found termination order: " ^ quote (Syntax.string_of_term ctxt relation))
-
- in (* 4: proof reconstruction *)
- st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)])
- THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
- THEN (rtac @{thm "wf_empty"} 1)
- THEN EVERY (map prove_row clean_table))
- end
-
-fun lexicographic_order_tac ctxt =
- TRY (FundefCommon.apply_termination_rule ctxt 1)
- THEN lex_order_tac ctxt (auto_tac (local_clasimpset_of ctxt addsimps2 FundefCommon.TerminationSimps.get ctxt))
-
-val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac
-
-val setup =
- Method.setup @{binding lexicographic_order}
- (Method.sections clasimp_modifiers >> (K lexicographic_order))
- "termination prover for lexicographic orderings"
- #> Context.theory_map (FundefCommon.set_termination_prover lexicographic_order)
-
-end;
-
--- a/src/HOL/Tools/function_package/measure_functions.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,59 +0,0 @@
-(* Title: HOL/Tools/function_package/measure_functions.ML
- ID: $Id$
- Author: Alexander Krauss, TU Muenchen
-
-Measure functions, generated heuristically
-*)
-
-signature MEASURE_FUNCTIONS =
-sig
-
- val get_measure_functions : Proof.context -> typ -> term list
- val setup : theory -> theory
-
-end
-
-structure MeasureFunctions : MEASURE_FUNCTIONS =
-struct
-
-(** User-declared size functions **)
-structure MeasureHeuristicRules = NamedThmsFun(
- val name = "measure_function"
- val description = "Rules that guide the heuristic generation of measure functions"
-);
-
-fun mk_is_measures t = Const (@{const_name "is_measure"}, fastype_of t --> HOLogic.boolT) $ t
-
-fun find_measures ctxt T =
- DEPTH_SOLVE (resolve_tac (MeasureHeuristicRules.get ctxt) 1)
- (HOLogic.mk_Trueprop (mk_is_measures (Var (("f",0), T --> HOLogic.natT)))
- |> cterm_of (ProofContext.theory_of ctxt) |> Goal.init)
- |> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
- |> Seq.list_of
-
-
-(** Generating Measure Functions **)
-
-fun constant_0 T = Abs ("x", T, HOLogic.zero)
-fun constant_1 T = Abs ("x", T, HOLogic.Suc_zero)
-
-fun mk_funorder_funs (Type ("+", [fT, sT])) =
- map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT m (constant_0 sT)) (mk_funorder_funs fT)
- @ map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT (constant_0 fT) m) (mk_funorder_funs sT)
- | mk_funorder_funs T = [ constant_1 T ]
-
-fun mk_ext_base_funs ctxt (Type("+", [fT, sT])) =
- map_product (SumTree.mk_sumcase fT sT HOLogic.natT)
- (mk_ext_base_funs ctxt fT) (mk_ext_base_funs ctxt sT)
- | mk_ext_base_funs ctxt T = find_measures ctxt T
-
-fun mk_all_measure_funs ctxt (T as Type ("+", _)) =
- mk_ext_base_funs ctxt T @ mk_funorder_funs T
- | mk_all_measure_funs ctxt T = find_measures ctxt T
-
-val get_measure_functions = mk_all_measure_funs
-
-val setup = MeasureHeuristicRules.setup
-
-end
-
--- a/src/HOL/Tools/function_package/mutual.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,314 +0,0 @@
-(* Title: HOL/Tools/function_package/mutual.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Tools for mutual recursive definitions.
-*)
-
-signature FUNDEF_MUTUAL =
-sig
-
- val prepare_fundef_mutual : FundefCommon.fundef_config
- -> string (* defname *)
- -> ((string * typ) * mixfix) list
- -> term list
- -> local_theory
- -> ((thm (* goalstate *)
- * (thm -> FundefCommon.fundef_result) (* proof continuation *)
- ) * local_theory)
-
-end
-
-
-structure FundefMutual: FUNDEF_MUTUAL =
-struct
-
-open FundefLib
-open FundefCommon
-
-
-
-
-type qgar = string * (string * typ) list * term list * term list * term
-
-fun name_of_fqgar ((f, _, _, _, _): qgar) = f
-
-datatype mutual_part =
- MutualPart of
- {
- i : int,
- i' : int,
- fvar : string * typ,
- cargTs: typ list,
- f_def: term,
-
- f: term option,
- f_defthm : thm option
- }
-
-
-datatype mutual_info =
- Mutual of
- {
- n : int,
- n' : int,
- fsum_var : string * typ,
-
- ST: typ,
- RST: typ,
-
- parts: mutual_part list,
- fqgars: qgar list,
- qglrs: ((string * typ) list * term list * term * term) list,
-
- fsum : term option
- }
-
-fun mutual_induct_Pnames n =
- if n < 5 then fst (chop n ["P","Q","R","S"])
- else map (fn i => "P" ^ string_of_int i) (1 upto n)
-
-fun get_part fname =
- the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
-
-(* FIXME *)
-fun mk_prod_abs e (t1, t2) =
- let
- val bTs = rev (map snd e)
- val T1 = fastype_of1 (bTs, t1)
- val T2 = fastype_of1 (bTs, t2)
- in
- HOLogic.pair_const T1 T2 $ t1 $ t2
- end;
-
-
-fun analyze_eqs ctxt defname fs eqs =
- let
- val num = length fs
- val fnames = map fst fs
- val fqgars = map (split_def ctxt) eqs
- val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
- |> AList.lookup (op =) #> the
-
- fun curried_types (fname, fT) =
- let
- val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
- in
- (caTs, uaTs ---> body_type fT)
- end
-
- val (caTss, resultTs) = split_list (map curried_types fs)
- val argTs = map (foldr1 HOLogic.mk_prodT) caTss
-
- val dresultTs = distinct (Type.eq_type Vartab.empty) resultTs
- val n' = length dresultTs
-
- val RST = BalancedTree.make (uncurry SumTree.mk_sumT) dresultTs
- val ST = BalancedTree.make (uncurry SumTree.mk_sumT) argTs
-
- val fsum_type = ST --> RST
-
- val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
- val fsum_var = (fsum_var_name, fsum_type)
-
- fun define (fvar as (n, T)) caTs resultT i =
- let
- val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
- val i' = find_index (fn Ta => Type.eq_type Vartab.empty (Ta, resultT)) dresultTs + 1
-
- val f_exp = SumTree.mk_proj RST n' i' (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
- val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
-
- val rew = (n, fold_rev lambda vars f_exp)
- in
- (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
- end
-
- val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
-
- fun convert_eqs (f, qs, gs, args, rhs) =
- let
- val MutualPart {i, i', ...} = get_part f parts
- in
- (qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
- SumTree.mk_inj RST n' i' (replace_frees rews rhs)
- |> Envir.beta_norm)
- end
-
- val qglrs = map convert_eqs fqgars
- in
- Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST,
- parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
- end
-
-
-
-
-fun define_projections fixes mutual fsum lthy =
- let
- fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
- let
- val ((f, (_, f_defthm)), lthy') =
- LocalTheory.define Thm.internalK ((Binding.name fname, mixfix),
- ((Binding.name (fname ^ "_def"), []), Term.subst_bound (fsum, f_def)))
- lthy
- in
- (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
- f=SOME f, f_defthm=SOME f_defthm },
- lthy')
- end
-
- val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
- val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
- in
- (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
- fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
- lthy')
- end
-
-
-fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
- let
- val thy = ProofContext.theory_of ctxt
-
- val oqnames = map fst pre_qs
- val (qs, ctxt') = Variable.variant_fixes oqnames ctxt
- |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
-
- fun inst t = subst_bounds (rev qs, t)
- val gs = map inst pre_gs
- val args = map inst pre_args
- val rhs = inst pre_rhs
-
- val cqs = map (cterm_of thy) qs
- val ags = map (assume o cterm_of thy) gs
-
- val import = fold forall_elim cqs
- #> fold Thm.elim_implies ags
-
- val export = fold_rev (implies_intr o cprop_of) ags
- #> fold_rev forall_intr_rename (oqnames ~~ cqs)
- in
- F ctxt (f, qs, gs, args, rhs) import export
- end
-
-fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs) import (export : thm -> thm) sum_psimp_eq =
- let
- val (MutualPart {f=SOME f, f_defthm=SOME f_def, ...}) = get_part fname parts
-
- val psimp = import sum_psimp_eq
- val (simp, restore_cond) = case cprems_of psimp of
- [] => (psimp, I)
- | [cond] => (implies_elim psimp (assume cond), implies_intr cond)
- | _ => sys_error "Too many conditions"
- in
- Goal.prove ctxt [] []
- (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
- (fn _ => (LocalDefs.unfold_tac ctxt all_orig_fdefs)
- THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
- THEN (simp_tac (local_simpset_of ctxt addsimps SumTree.proj_in_rules)) 1)
- |> restore_cond
- |> export
- end
-
-
-(* FIXME HACK *)
-fun mk_applied_form ctxt caTs thm =
- let
- val thy = ProofContext.theory_of ctxt
- val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
- in
- fold (fn x => fn thm => combination thm (reflexive x)) xs thm
- |> Conv.fconv_rule (Thm.beta_conversion true)
- |> fold_rev forall_intr xs
- |> Thm.forall_elim_vars 0
- end
-
-
-fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, RST, parts, ...}) =
- let
- val cert = cterm_of (ProofContext.theory_of lthy)
- val newPs = map2 (fn Pname => fn MutualPart {cargTs, ...} =>
- Free (Pname, cargTs ---> HOLogic.boolT))
- (mutual_induct_Pnames (length parts))
- parts
-
- fun mk_P (MutualPart {cargTs, ...}) P =
- let
- val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
- val atup = foldr1 HOLogic.mk_prod avars
- in
- tupled_lambda atup (list_comb (P, avars))
- end
-
- val Ps = map2 mk_P parts newPs
- val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
-
- val induct_inst =
- forall_elim (cert case_exp) induct
- |> full_simplify SumTree.sumcase_split_ss
- |> full_simplify (HOL_basic_ss addsimps all_f_defs)
-
- fun project rule (MutualPart {cargTs, i, ...}) k =
- let
- val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
- val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
- in
- (rule
- |> forall_elim (cert inj)
- |> full_simplify SumTree.sumcase_split_ss
- |> fold_rev (forall_intr o cert) (afs @ newPs),
- k + length cargTs)
- end
- in
- fst (fold_map (project induct_inst) parts 0)
- end
-
-
-fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
- let
- val result = inner_cont proof
- val FundefResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
- termination,domintros} = result
-
- val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
- (mk_applied_form lthy cargTs (symmetric f_def), f))
- parts
- |> split_list
-
- val all_orig_fdefs = map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
-
- fun mk_mpsimp fqgar sum_psimp =
- in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
-
- val rew_ss = HOL_basic_ss addsimps all_f_defs
- val mpsimps = map2 mk_mpsimp fqgars psimps
- val mtrsimps = map_option (map2 mk_mpsimp fqgars) trsimps
- val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
- val mtermination = full_simplify rew_ss termination
- val mdomintros = map_option (map (full_simplify rew_ss)) domintros
- in
- FundefResult { fs=fs, G=G, R=R,
- psimps=mpsimps, simple_pinducts=minducts,
- cases=cases, termination=mtermination,
- domintros=mdomintros,
- trsimps=mtrsimps}
- end
-
-fun prepare_fundef_mutual config defname fixes eqss lthy =
- let
- val mutual = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
- val Mutual {fsum_var=(n, T), qglrs, ...} = mutual
-
- val ((fsum, goalstate, cont), lthy') =
- FundefCore.prepare_fundef config defname [((n, T), NoSyn)] qglrs lthy
-
- val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
-
- val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual'
- in
- ((goalstate, mutual_cont), lthy'')
- end
-
-
-end
--- a/src/HOL/Tools/function_package/pattern_split.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,138 +0,0 @@
-(* Title: HOL/Tools/function_package/pattern_split.ML
- ID: $Id$
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-
-Automatic splitting of overlapping constructor patterns. This is a preprocessing step which
-turns a specification with overlaps into an overlap-free specification.
-
-*)
-
-signature FUNDEF_SPLIT =
-sig
- val split_some_equations :
- Proof.context -> (bool * term) list -> term list list
-
- val split_all_equations :
- Proof.context -> term list -> term list list
-end
-
-structure FundefSplit : FUNDEF_SPLIT =
-struct
-
-open FundefLib
-
-(* We use proof context for the variable management *)
-(* FIXME: no __ *)
-
-fun new_var ctx vs T =
- let
- val [v] = Variable.variant_frees ctx vs [("v", T)]
- in
- (Free v :: vs, Free v)
- end
-
-fun saturate ctx vs t =
- fold (fn T => fn (vs, t) => new_var ctx vs T |> apsnd (curry op $ t))
- (binder_types (fastype_of t)) (vs, t)
-
-
-(* This is copied from "fundef_datatype.ML" *)
-fun inst_constrs_of thy (T as Type (name, _)) =
- map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
- (the (DatatypePackage.get_datatype_constrs thy name))
- | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
-
-
-
-
-fun join ((vs1,sub1), (vs2,sub2)) = (merge (op aconv) (vs1,vs2), sub1 @ sub2)
-fun join_product (xs, ys) = map_product (curry join) xs ys
-
-fun join_list [] = []
- | join_list xs = foldr1 (join_product) xs
-
-
-exception DISJ
-
-fun pattern_subtract_subst ctx vs t t' =
- let
- exception DISJ
- fun pattern_subtract_subst_aux vs _ (Free v2) = []
- | pattern_subtract_subst_aux vs (v as (Free (_, T))) t' =
- let
- fun foo constr =
- let
- val (vs', t) = saturate ctx vs constr
- val substs = pattern_subtract_subst ctx vs' t t'
- in
- map (fn (vs, subst) => (vs, (v,t)::subst)) substs
- end
- in
- flat (map foo (inst_constrs_of (ProofContext.theory_of ctx) T))
- end
- | pattern_subtract_subst_aux vs t t' =
- let
- val (C, ps) = strip_comb t
- val (C', qs) = strip_comb t'
- in
- if C = C'
- then flat (map2 (pattern_subtract_subst_aux vs) ps qs)
- else raise DISJ
- end
- in
- pattern_subtract_subst_aux vs t t'
- handle DISJ => [(vs, [])]
- end
-
-
-(* p - q *)
-fun pattern_subtract ctx eq2 eq1 =
- let
- val thy = ProofContext.theory_of ctx
-
- val (vs, feq1 as (_ $ (_ $ lhs1 $ _))) = dest_all_all eq1
- val (_, _ $ (_ $ lhs2 $ _)) = dest_all_all eq2
-
- val substs = pattern_subtract_subst ctx vs lhs1 lhs2
-
- fun instantiate (vs', sigma) =
- let
- val t = Pattern.rewrite_term thy sigma [] feq1
- in
- fold_rev Logic.all (map Free (frees_in_term ctx t) inter vs') t
- end
- in
- map instantiate substs
- end
-
-
-(* ps - p' *)
-fun pattern_subtract_from_many ctx p'=
- flat o map (pattern_subtract ctx p')
-
-(* in reverse order *)
-fun pattern_subtract_many ctx ps' =
- fold_rev (pattern_subtract_from_many ctx) ps'
-
-
-
-fun split_some_equations ctx eqns =
- let
- fun split_aux prev [] = []
- | split_aux prev ((true, eq) :: es) = pattern_subtract_many ctx prev [eq]
- :: split_aux (eq :: prev) es
- | split_aux prev ((false, eq) :: es) = [eq]
- :: split_aux (eq :: prev) es
- in
- split_aux [] eqns
- end
-
-fun split_all_equations ctx =
- split_some_equations ctx o map (pair true)
-
-
-
-
-end
--- a/src/HOL/Tools/function_package/scnp_reconstruct.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,429 +0,0 @@
-(* Title: HOL/Tools/function_package/scnp_reconstruct.ML
- Author: Armin Heller, TU Muenchen
- Author: Alexander Krauss, TU Muenchen
-
-Proof reconstruction for SCNP
-*)
-
-signature SCNP_RECONSTRUCT =
-sig
-
- val sizechange_tac : Proof.context -> tactic -> tactic
-
- val decomp_scnp : ScnpSolve.label list -> Proof.context -> Proof.method
-
- val setup : theory -> theory
-
- datatype multiset_setup =
- Multiset of
- {
- msetT : typ -> typ,
- mk_mset : typ -> term list -> term,
- mset_regroup_conv : int list -> conv,
- mset_member_tac : int -> int -> tactic,
- mset_nonempty_tac : int -> tactic,
- mset_pwleq_tac : int -> tactic,
- set_of_simps : thm list,
- smsI' : thm,
- wmsI2'' : thm,
- wmsI1 : thm,
- reduction_pair : thm
- }
-
-
- val multiset_setup : multiset_setup -> theory -> theory
-
-end
-
-structure ScnpReconstruct : SCNP_RECONSTRUCT =
-struct
-
-val PROFILE = FundefCommon.PROFILE
-fun TRACE x = if ! FundefCommon.profile then Output.tracing x else ()
-
-open ScnpSolve
-
-val natT = HOLogic.natT
-val nat_pairT = HOLogic.mk_prodT (natT, natT)
-
-(* Theory dependencies *)
-
-datatype multiset_setup =
- Multiset of
- {
- msetT : typ -> typ,
- mk_mset : typ -> term list -> term,
- mset_regroup_conv : int list -> conv,
- mset_member_tac : int -> int -> tactic,
- mset_nonempty_tac : int -> tactic,
- mset_pwleq_tac : int -> tactic,
- set_of_simps : thm list,
- smsI' : thm,
- wmsI2'' : thm,
- wmsI1 : thm,
- reduction_pair : thm
- }
-
-structure MultisetSetup = TheoryDataFun
-(
- type T = multiset_setup option
- val empty = NONE
- val copy = I;
- val extend = I;
- fun merge _ (v1, v2) = if is_some v2 then v2 else v1
-)
-
-val multiset_setup = MultisetSetup.put o SOME
-
-fun undef x = error "undef"
-fun get_multiset_setup thy = MultisetSetup.get thy
- |> the_default (Multiset
-{ msetT = undef, mk_mset=undef,
- mset_regroup_conv=undef, mset_member_tac = undef,
- mset_nonempty_tac = undef, mset_pwleq_tac = undef,
- set_of_simps = [],reduction_pair = refl,
- smsI'=refl, wmsI2''=refl, wmsI1=refl })
-
-fun order_rpair _ MAX = @{thm max_rpair_set}
- | order_rpair msrp MS = msrp
- | order_rpair _ MIN = @{thm min_rpair_set}
-
-fun ord_intros_max true =
- (@{thm smax_emptyI}, @{thm smax_insertI})
- | ord_intros_max false =
- (@{thm wmax_emptyI}, @{thm wmax_insertI})
-fun ord_intros_min true =
- (@{thm smin_emptyI}, @{thm smin_insertI})
- | ord_intros_min false =
- (@{thm wmin_emptyI}, @{thm wmin_insertI})
-
-fun gen_probl D cs =
- let
- val n = Termination.get_num_points D
- val arity = length o Termination.get_measures D
- fun measure p i = nth (Termination.get_measures D p) i
-
- fun mk_graph c =
- let
- val (_, p, _, q, _, _) = Termination.dest_call D c
-
- fun add_edge i j =
- case Termination.get_descent D c (measure p i) (measure q j)
- of SOME (Termination.Less _) => cons (i, GTR, j)
- | SOME (Termination.LessEq _) => cons (i, GEQ, j)
- | _ => I
-
- val edges =
- fold_product add_edge (0 upto arity p - 1) (0 upto arity q - 1) []
- in
- G (p, q, edges)
- end
- in
- GP (map arity (0 upto n - 1), map mk_graph cs)
- end
-
-(* General reduction pair application *)
-fun rem_inv_img ctxt =
- let
- val unfold_tac = LocalDefs.unfold_tac ctxt
- in
- rtac @{thm subsetI} 1
- THEN etac @{thm CollectE} 1
- THEN REPEAT (etac @{thm exE} 1)
- THEN unfold_tac @{thms inv_image_def}
- THEN rtac @{thm CollectI} 1
- THEN etac @{thm conjE} 1
- THEN etac @{thm ssubst} 1
- THEN unfold_tac (@{thms split_conv} @ @{thms triv_forall_equality}
- @ @{thms sum.cases})
- end
-
-(* Sets *)
-
-val setT = HOLogic.mk_setT
-
-fun set_member_tac m i =
- if m = 0 then rtac @{thm insertI1} i
- else rtac @{thm insertI2} i THEN set_member_tac (m - 1) i
-
-val set_nonempty_tac = rtac @{thm insert_not_empty}
-
-fun set_finite_tac i =
- rtac @{thm finite.emptyI} i
- ORELSE (rtac @{thm finite.insertI} i THEN (fn st => set_finite_tac i st))
-
-
-(* Reconstruction *)
-
-fun reconstruct_tac ctxt D cs (gp as GP (_, gs)) certificate =
- let
- val thy = ProofContext.theory_of ctxt
- val Multiset
- { msetT, mk_mset,
- mset_regroup_conv, mset_member_tac,
- mset_nonempty_tac, mset_pwleq_tac, set_of_simps,
- smsI', wmsI2'', wmsI1, reduction_pair=ms_rp }
- = get_multiset_setup thy
-
- fun measure_fn p = nth (Termination.get_measures D p)
-
- fun get_desc_thm cidx m1 m2 bStrict =
- case Termination.get_descent D (nth cs cidx) m1 m2
- of SOME (Termination.Less thm) =>
- if bStrict then thm
- else (thm COMP (Thm.lift_rule (cprop_of thm) @{thm less_imp_le}))
- | SOME (Termination.LessEq (thm, _)) =>
- if not bStrict then thm
- else sys_error "get_desc_thm"
- | _ => sys_error "get_desc_thm"
-
- val (label, lev, sl, covering) = certificate
-
- fun prove_lev strict g =
- let
- val G (p, q, el) = nth gs g
-
- fun less_proof strict (j, b) (i, a) =
- let
- val tag_flag = b < a orelse (not strict andalso b <= a)
-
- val stored_thm =
- get_desc_thm g (measure_fn p i) (measure_fn q j)
- (not tag_flag)
- |> Conv.fconv_rule (Thm.beta_conversion true)
-
- val rule = if strict
- then if b < a then @{thm pair_lessI2} else @{thm pair_lessI1}
- else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
- in
- rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
- THEN (if tag_flag then Arith_Data.verbose_arith_tac ctxt 1 else all_tac)
- end
-
- fun steps_tac MAX strict lq lp =
- let
- val (empty, step) = ord_intros_max strict
- in
- if length lq = 0
- then rtac empty 1 THEN set_finite_tac 1
- THEN (if strict then set_nonempty_tac 1 else all_tac)
- else
- let
- val (j, b) :: rest = lq
- val (i, a) = the (covering g strict j)
- fun choose xs = set_member_tac (Library.find_index (curry op = (i, a)) xs) 1
- val solve_tac = choose lp THEN less_proof strict (j, b) (i, a)
- in
- rtac step 1 THEN solve_tac THEN steps_tac MAX strict rest lp
- end
- end
- | steps_tac MIN strict lq lp =
- let
- val (empty, step) = ord_intros_min strict
- in
- if length lp = 0
- then rtac empty 1
- THEN (if strict then set_nonempty_tac 1 else all_tac)
- else
- let
- val (i, a) :: rest = lp
- val (j, b) = the (covering g strict i)
- fun choose xs = set_member_tac (Library.find_index (curry op = (j, b)) xs) 1
- val solve_tac = choose lq THEN less_proof strict (j, b) (i, a)
- in
- rtac step 1 THEN solve_tac THEN steps_tac MIN strict lq rest
- end
- end
- | steps_tac MS strict lq lp =
- let
- fun get_str_cover (j, b) =
- if is_some (covering g true j) then SOME (j, b) else NONE
- fun get_wk_cover (j, b) = the (covering g false j)
-
- val qs = lq \\ map_filter get_str_cover lq
- val ps = map get_wk_cover qs
-
- fun indices xs ys = map (fn y => Library.find_index (curry op = y) xs) ys
- val iqs = indices lq qs
- val ips = indices lp ps
-
- local open Conv in
- fun t_conv a C =
- params_conv ~1 (K ((concl_conv ~1 o arg_conv o arg1_conv o a) C)) ctxt
- val goal_rewrite =
- t_conv arg1_conv (mset_regroup_conv iqs)
- then_conv t_conv arg_conv (mset_regroup_conv ips)
- end
- in
- CONVERSION goal_rewrite 1
- THEN (if strict then rtac smsI' 1
- else if qs = lq then rtac wmsI2'' 1
- else rtac wmsI1 1)
- THEN mset_pwleq_tac 1
- THEN EVERY (map2 (less_proof false) qs ps)
- THEN (if strict orelse qs <> lq
- then LocalDefs.unfold_tac ctxt set_of_simps
- THEN steps_tac MAX true (lq \\ qs) (lp \\ ps)
- else all_tac)
- end
- in
- rem_inv_img ctxt
- THEN steps_tac label strict (nth lev q) (nth lev p)
- end
-
- val (mk_set, setT) = if label = MS then (mk_mset, msetT) else (HOLogic.mk_set, setT)
-
- fun tag_pair p (i, tag) =
- HOLogic.pair_const natT natT $
- (measure_fn p i $ Bound 0) $ HOLogic.mk_number natT tag
-
- fun pt_lev (p, lm) = Abs ("x", Termination.get_types D p,
- mk_set nat_pairT (map (tag_pair p) lm))
-
- val level_mapping =
- map_index pt_lev lev
- |> Termination.mk_sumcases D (setT nat_pairT)
- |> cterm_of thy
- in
- PROFILE "Proof Reconstruction"
- (CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv sl))) 1
- THEN (rtac @{thm reduction_pair_lemma} 1)
- THEN (rtac @{thm rp_inv_image_rp} 1)
- THEN (rtac (order_rpair ms_rp label) 1)
- THEN PRIMITIVE (instantiate' [] [SOME level_mapping])
- THEN unfold_tac @{thms rp_inv_image_def} (local_simpset_of ctxt)
- THEN LocalDefs.unfold_tac ctxt
- (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv})
- THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}]))
- THEN EVERY (map (prove_lev true) sl)
- THEN EVERY (map (prove_lev false) ((0 upto length cs - 1) \\ sl)))
- end
-
-
-
-local open Termination in
-fun print_cell (SOME (Less _)) = "<"
- | print_cell (SOME (LessEq _)) = "\<le>"
- | print_cell (SOME (None _)) = "-"
- | print_cell (SOME (False _)) = "-"
- | print_cell (NONE) = "?"
-
-fun print_error ctxt D = CALLS (fn (cs, i) =>
- let
- val np = get_num_points D
- val ms = map (get_measures D) (0 upto np - 1)
- val tys = map (get_types D) (0 upto np - 1)
- fun index xs = (1 upto length xs) ~~ xs
- fun outp s t f xs = map (fn (x, y) => s ^ Int.toString x ^ t ^ f y ^ "\n") xs
- val ims = index (map index ms)
- val _ = Output.tracing (concat (outp "fn #" ":\n" (concat o outp "\tmeasure #" ": " (Syntax.string_of_term ctxt)) ims))
- fun print_call (k, c) =
- let
- val (_, p, _, q, _, _) = dest_call D c
- val _ = Output.tracing ("call table for call #" ^ Int.toString k ^ ": fn " ^
- Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1))
- val caller_ms = nth ms p
- val callee_ms = nth ms q
- val entries = map (fn x => map (pair x) (callee_ms)) (caller_ms)
- fun print_ln (i : int, l) = concat (Int.toString i :: " " :: map (enclose " " " " o print_cell o (uncurry (get_descent D c))) l)
- val _ = Output.tracing (concat (Int.toString (p + 1) ^ "|" ^ Int.toString (q + 1) ^
- " " :: map (enclose " " " " o Int.toString) (1 upto length callee_ms)) ^ "\n"
- ^ cat_lines (map print_ln ((1 upto (length entries)) ~~ entries)))
- in
- true
- end
- fun list_call (k, c) =
- let
- val (_, p, _, q, _, _) = dest_call D c
- val _ = Output.tracing ("call #" ^ (Int.toString k) ^ ": fn " ^
- Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1) ^ "\n" ^
- (Syntax.string_of_term ctxt c))
- in true end
- val _ = forall list_call ((1 upto length cs) ~~ cs)
- val _ = forall print_call ((1 upto length cs) ~~ cs)
- in
- all_tac
- end)
-end
-
-
-fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
- let
- val ms_configured = is_some (MultisetSetup.get (ProofContext.theory_of ctxt))
- val orders' = if ms_configured then orders
- else filter_out (curry op = MS) orders
- val gp = gen_probl D cs
-(* val _ = TRACE ("SCNP instance: " ^ makestring gp)*)
- val certificate = generate_certificate use_tags orders' gp
-(* val _ = TRACE ("Certificate: " ^ makestring certificate)*)
-
- in
- case certificate
- of NONE => err_cont D i
- | SOME cert =>
- SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i
- THEN (rtac @{thm wf_empty} i ORELSE cont D i)
- end)
-
-fun gen_decomp_scnp_tac orders autom_tac ctxt err_cont =
- let
- open Termination
- val derive_diag = Descent.derive_diag ctxt autom_tac
- val derive_all = Descent.derive_all ctxt autom_tac
- val decompose = Decompose.decompose_tac ctxt autom_tac
- val scnp_no_tags = single_scnp_tac false orders ctxt
- val scnp_full = single_scnp_tac true orders ctxt
-
- fun first_round c e =
- derive_diag (REPEAT scnp_no_tags c e)
-
- val second_round =
- REPEAT (fn c => fn e => decompose (scnp_no_tags c c) e)
-
- val third_round =
- derive_all oo
- REPEAT (fn c => fn e =>
- scnp_full (decompose c c) e)
-
- fun Then s1 s2 c e = s1 (s2 c c) (s2 c e)
-
- val strategy = Then (Then first_round second_round) third_round
-
- in
- TERMINATION ctxt (strategy err_cont err_cont)
- end
-
-fun gen_sizechange_tac orders autom_tac ctxt err_cont =
- TRY (FundefCommon.apply_termination_rule ctxt 1)
- THEN TRY (Termination.wf_union_tac ctxt)
- THEN
- (rtac @{thm wf_empty} 1
- ORELSE gen_decomp_scnp_tac orders autom_tac ctxt err_cont 1)
-
-fun sizechange_tac ctxt autom_tac =
- gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt (K (K all_tac))
-
-fun decomp_scnp orders ctxt =
- let
- val extra_simps = FundefCommon.TerminationSimps.get ctxt
- val autom_tac = auto_tac (local_clasimpset_of ctxt addsimps2 extra_simps)
- in
- SIMPLE_METHOD
- (gen_sizechange_tac orders autom_tac ctxt (print_error ctxt))
- end
-
-
-(* Method setup *)
-
-val orders =
- Scan.repeat1
- ((Args.$$$ "max" >> K MAX) ||
- (Args.$$$ "min" >> K MIN) ||
- (Args.$$$ "ms" >> K MS))
- || Scan.succeed [MAX, MS, MIN]
-
-val setup = Method.setup @{binding sizechange}
- (Scan.lift orders --| Method.sections clasimp_modifiers >> decomp_scnp)
- "termination prover with graph decomposition and the NP subset of size change termination"
-
-end
--- a/src/HOL/Tools/function_package/scnp_solve.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,257 +0,0 @@
-(* Title: HOL/Tools/function_package/scnp_solve.ML
- Author: Armin Heller, TU Muenchen
- Author: Alexander Krauss, TU Muenchen
-
-Generate certificates for SCNP using a SAT solver
-*)
-
-
-signature SCNP_SOLVE =
-sig
-
- datatype edge = GTR | GEQ
- datatype graph = G of int * int * (int * edge * int) list
- datatype graph_problem = GP of int list * graph list
-
- datatype label = MIN | MAX | MS
-
- type certificate =
- label (* which order *)
- * (int * int) list list (* (multi)sets *)
- * int list (* strictly ordered calls *)
- * (int -> bool -> int -> (int * int) option) (* covering function *)
-
- val generate_certificate : bool -> label list -> graph_problem -> certificate option
-
- val solver : string ref
-end
-
-structure ScnpSolve : SCNP_SOLVE =
-struct
-
-(** Graph problems **)
-
-datatype edge = GTR | GEQ ;
-datatype graph = G of int * int * (int * edge * int) list ;
-datatype graph_problem = GP of int list * graph list ;
-
-datatype label = MIN | MAX | MS ;
-type certificate =
- label
- * (int * int) list list
- * int list
- * (int -> bool -> int -> (int * int) option)
-
-fun graph_at (GP (_, gs), i) = nth gs i ;
-fun num_prog_pts (GP (arities, _)) = length arities ;
-fun num_graphs (GP (_, gs)) = length gs ;
-fun arity (GP (arities, gl)) i = nth arities i ;
-fun ndigits (GP (arities, _)) = IntInf.log2 (List.foldl (op +) 0 arities) + 1
-
-
-(** Propositional formulas **)
-
-val Not = PropLogic.Not and And = PropLogic.And and Or = PropLogic.Or
-val BoolVar = PropLogic.BoolVar
-fun Implies (p, q) = Or (Not p, q)
-fun Equiv (p, q) = And (Implies (p, q), Implies (q, p))
-val all = PropLogic.all
-
-(* finite indexed quantifiers:
-
-iforall n f <==> /\
- / \ f i
- 0<=i<n
- *)
-fun iforall n f = all (map f (0 upto n - 1))
-fun iexists n f = PropLogic.exists (map f (0 upto n - 1))
-fun iforall2 n m f = all (map_product f (0 upto n - 1) (0 upto m - 1))
-
-fun the_one var n x = all (var x :: map (Not o var) ((0 upto n - 1) \\ [x]))
-fun exactly_one n f = iexists n (the_one f n)
-
-(* SAT solving *)
-val solver = ref "auto";
-fun sat_solver x =
- FundefCommon.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
-
-(* "Virtual constructors" for various propositional variables *)
-fun var_constrs (gp as GP (arities, gl)) =
- let
- val n = Int.max (num_graphs gp, num_prog_pts gp)
- val k = List.foldl Int.max 1 arities
-
- (* Injective, provided a < 8, x < n, and i < k. *)
- fun prod a x i j = ((j * k + i) * n + x) * 8 + a + 1
-
- fun ES (g, i, j) = BoolVar (prod 0 g i j)
- fun EW (g, i, j) = BoolVar (prod 1 g i j)
- fun WEAK g = BoolVar (prod 2 g 0 0)
- fun STRICT g = BoolVar (prod 3 g 0 0)
- fun P (p, i) = BoolVar (prod 4 p i 0)
- fun GAM (g, i, j)= BoolVar (prod 5 g i j)
- fun EPS (g, i) = BoolVar (prod 6 g i 0)
- fun TAG (p, i) b = BoolVar (prod 7 p i b)
- in
- (ES,EW,WEAK,STRICT,P,GAM,EPS,TAG)
- end
-
-
-fun graph_info gp g =
- let
- val G (p, q, edgs) = graph_at (gp, g)
- in
- (g, p, q, arity gp p, arity gp q, edgs)
- end
-
-
-(* Order-independent part of encoding *)
-
-fun encode_graphs bits gp =
- let
- val ng = num_graphs gp
- val (ES,EW,_,_,_,_,_,TAG) = var_constrs gp
-
- fun encode_constraint_strict 0 (x, y) = PropLogic.False
- | encode_constraint_strict k (x, y) =
- Or (And (TAG x (k - 1), Not (TAG y (k - 1))),
- And (Equiv (TAG x (k - 1), TAG y (k - 1)),
- encode_constraint_strict (k - 1) (x, y)))
-
- fun encode_constraint_weak k (x, y) =
- Or (encode_constraint_strict k (x, y),
- iforall k (fn i => Equiv (TAG x i, TAG y i)))
-
- fun encode_graph (g, p, q, n, m, edges) =
- let
- fun encode_edge i j =
- if exists (fn x => x = (i, GTR, j)) edges then
- And (ES (g, i, j), EW (g, i, j))
- else if not (exists (fn x => x = (i, GEQ, j)) edges) then
- And (Not (ES (g, i, j)), Not (EW (g, i, j)))
- else
- And (
- Equiv (ES (g, i, j),
- encode_constraint_strict bits ((p, i), (q, j))),
- Equiv (EW (g, i, j),
- encode_constraint_weak bits ((p, i), (q, j))))
- in
- iforall2 n m encode_edge
- end
- in
- iforall ng (encode_graph o graph_info gp)
- end
-
-
-(* Order-specific part of encoding *)
-
-fun encode bits gp mu =
- let
- val ng = num_graphs gp
- val (ES,EW,WEAK,STRICT,P,GAM,EPS,_) = var_constrs gp
-
- fun encode_graph MAX (g, p, q, n, m, _) =
- And (
- Equiv (WEAK g,
- iforall m (fn j =>
- Implies (P (q, j),
- iexists n (fn i =>
- And (P (p, i), EW (g, i, j)))))),
- Equiv (STRICT g,
- And (
- iforall m (fn j =>
- Implies (P (q, j),
- iexists n (fn i =>
- And (P (p, i), ES (g, i, j))))),
- iexists n (fn i => P (p, i)))))
- | encode_graph MIN (g, p, q, n, m, _) =
- And (
- Equiv (WEAK g,
- iforall n (fn i =>
- Implies (P (p, i),
- iexists m (fn j =>
- And (P (q, j), EW (g, i, j)))))),
- Equiv (STRICT g,
- And (
- iforall n (fn i =>
- Implies (P (p, i),
- iexists m (fn j =>
- And (P (q, j), ES (g, i, j))))),
- iexists m (fn j => P (q, j)))))
- | encode_graph MS (g, p, q, n, m, _) =
- all [
- Equiv (WEAK g,
- iforall m (fn j =>
- Implies (P (q, j),
- iexists n (fn i => GAM (g, i, j))))),
- Equiv (STRICT g,
- iexists n (fn i =>
- And (P (p, i), Not (EPS (g, i))))),
- iforall2 n m (fn i => fn j =>
- Implies (GAM (g, i, j),
- all [
- P (p, i),
- P (q, j),
- EW (g, i, j),
- Equiv (Not (EPS (g, i)), ES (g, i, j))])),
- iforall n (fn i =>
- Implies (And (P (p, i), EPS (g, i)),
- exactly_one m (fn j => GAM (g, i, j))))
- ]
- in
- all [
- encode_graphs bits gp,
- iforall ng (encode_graph mu o graph_info gp),
- iforall ng (fn x => WEAK x),
- iexists ng (fn x => STRICT x)
- ]
- end
-
-
-(*Generieren des level-mapping und diverser output*)
-fun mk_certificate bits label gp f =
- let
- val (ES,EW,WEAK,STRICT,P,GAM,EPS,TAG) = var_constrs gp
- fun assign (PropLogic.BoolVar v) = the_default false (f v)
- fun assignTag i j =
- (fold (fn x => fn y => 2 * y + (if assign (TAG (i, j) x) then 1 else 0))
- (bits - 1 downto 0) 0)
-
- val level_mapping =
- let fun prog_pt_mapping p =
- map_filter (fn x => if assign (P(p, x)) then SOME (x, assignTag p x) else NONE)
- (0 upto (arity gp p) - 1)
- in map prog_pt_mapping (0 upto num_prog_pts gp - 1) end
-
- val strict_list = filter (assign o STRICT) (0 upto num_graphs gp - 1)
-
- fun covering_pair g bStrict j =
- let
- val (_, p, q, n, m, _) = graph_info gp g
-
- fun cover MAX j = find_index (fn i => assign (P (p, i)) andalso assign (EW (g, i, j))) (0 upto n - 1)
- | cover MS k = find_index (fn i => assign (GAM (g, i, k))) (0 upto n - 1)
- | cover MIN i = find_index (fn j => assign (P (q, j)) andalso assign (EW (g, i, j))) (0 upto m - 1)
- fun cover_strict MAX j = find_index (fn i => assign (P (p, i)) andalso assign (ES (g, i, j))) (0 upto n - 1)
- | cover_strict MS k = find_index (fn i => assign (GAM (g, i, k)) andalso not (assign (EPS (g, i) ))) (0 upto n - 1)
- | cover_strict MIN i = find_index (fn j => assign (P (q, j)) andalso assign (ES (g, i, j))) (0 upto m - 1)
- val i = if bStrict then cover_strict label j else cover label j
- in
- find_first (fn x => fst x = i) (nth level_mapping (if label = MIN then q else p))
- end
- in
- (label, level_mapping, strict_list, covering_pair)
- end
-
-(*interface for the proof reconstruction*)
-fun generate_certificate use_tags labels gp =
- let
- val bits = if use_tags then ndigits gp else 0
- in
- get_first
- (fn l => case sat_solver (encode bits gp l) of
- SatSolver.SATISFIABLE f => SOME (mk_certificate bits l gp f)
- | _ => NONE)
- labels
- end
-end
--- a/src/HOL/Tools/function_package/size.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,242 +0,0 @@
-(* Title: HOL/Tools/function_package/size.ML
- Author: Stefan Berghofer, Florian Haftmann & Alexander Krauss, TU Muenchen
-
-Size functions for datatypes.
-*)
-
-signature SIZE =
-sig
- val size_thms: theory -> string -> thm list
- val setup: theory -> theory
-end;
-
-structure Size: SIZE =
-struct
-
-open DatatypeAux;
-
-structure SizeData = TheoryDataFun
-(
- type T = (string * thm list) Symtab.table;
- val empty = Symtab.empty;
- val copy = I
- val extend = I
- fun merge _ = Symtab.merge (K true);
-);
-
-val lookup_size = SizeData.get #> Symtab.lookup;
-
-fun plus (t1, t2) = Const ("HOL.plus_class.plus",
- HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2;
-
-fun size_of_type f g h (T as Type (s, Ts)) =
- (case f s of
- SOME t => SOME t
- | NONE => (case g s of
- SOME size_name =>
- SOME (list_comb (Const (size_name,
- map (fn U => U --> HOLogic.natT) Ts @ [T] ---> HOLogic.natT),
- map (size_of_type' f g h) Ts))
- | NONE => NONE))
- | size_of_type f g h (TFree (s, _)) = h s
-and size_of_type' f g h T = (case size_of_type f g h T of
- NONE => Abs ("x", T, HOLogic.zero)
- | SOME t => t);
-
-fun is_poly thy (DtType (name, dts)) =
- (case DatatypePackage.get_datatype thy name of
- NONE => false
- | SOME _ => exists (is_poly thy) dts)
- | is_poly _ _ = true;
-
-fun constrs_of thy name =
- let
- val {descr, index, ...} = DatatypePackage.the_datatype thy name
- val SOME (_, _, constrs) = AList.lookup op = descr index
- in constrs end;
-
-val app = curry (list_comb o swap);
-
-fun prove_size_thms (info : datatype_info) new_type_names thy =
- let
- val {descr, alt_names, sorts, rec_names, rec_rewrites, induction, ...} = info;
- val l = length new_type_names;
- val alt_names' = (case alt_names of
- NONE => replicate l NONE | SOME names => map SOME names);
- val descr' = List.take (descr, l);
- val (rec_names1, rec_names2) = chop l rec_names;
- val recTs = get_rec_types descr sorts;
- val (recTs1, recTs2) = chop l recTs;
- val (_, (_, paramdts, _)) :: _ = descr;
- val paramTs = map (typ_of_dtyp descr sorts) paramdts;
- val ((param_size_fs, param_size_fTs), f_names) = paramTs |>
- map (fn T as TFree (s, _) =>
- let
- val name = "f" ^ implode (tl (explode s));
- val U = T --> HOLogic.natT
- in
- (((s, Free (name, U)), U), name)
- end) |> split_list |>> split_list;
- val param_size = AList.lookup op = param_size_fs;
-
- val extra_rewrites = descr |> map (#1 o snd) |> distinct op = |>
- map_filter (Option.map snd o lookup_size thy) |> flat;
- val extra_size = Option.map fst o lookup_size thy;
-
- val (((size_names, size_fns), def_names), def_names') =
- recTs1 ~~ alt_names' |>
- map (fn (T as Type (s, _), optname) =>
- let
- val s' = the_default (Long_Name.base_name s) optname ^ "_size";
- val s'' = Sign.full_bname thy s'
- in
- (s'',
- (list_comb (Const (s'', param_size_fTs @ [T] ---> HOLogic.natT),
- map snd param_size_fs),
- (s' ^ "_def", s' ^ "_overloaded_def")))
- end) |> split_list ||>> split_list ||>> split_list;
- val overloaded_size_fns = map HOLogic.size_const recTs1;
-
- (* instantiation for primrec combinator *)
- fun size_of_constr b size_ofp ((_, cargs), (_, cargs')) =
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val k = length (filter is_rec_type cargs);
- val (ts, _, _) = fold_rev (fn ((dt, dt'), T) => fn (us, i, j) =>
- if is_rec_type dt then (Bound i :: us, i + 1, j + 1)
- else
- (if b andalso is_poly thy dt' then
- case size_of_type (K NONE) extra_size size_ofp T of
- NONE => us | SOME sz => sz $ Bound j :: us
- else us, i, j + 1))
- (cargs ~~ cargs' ~~ Ts) ([], 0, k);
- val t =
- if null ts andalso (not b orelse not (exists (is_poly thy) cargs'))
- then HOLogic.zero
- else foldl1 plus (ts @ [HOLogic.Suc_zero])
- in
- List.foldr (fn (T, t') => Abs ("x", T, t')) t (Ts @ replicate k HOLogic.natT)
- end;
-
- val fs = maps (fn (_, (name, _, constrs)) =>
- map (size_of_constr true param_size) (constrs ~~ constrs_of thy name)) descr;
- val fs' = maps (fn (n, (name, _, constrs)) =>
- map (size_of_constr (l <= n) (K NONE)) (constrs ~~ constrs_of thy name)) descr;
- val fTs = map fastype_of fs;
-
- val (rec_combs1, rec_combs2) = chop l (map (fn (T, rec_name) =>
- Const (rec_name, fTs @ [T] ---> HOLogic.natT))
- (recTs ~~ rec_names));
-
- fun define_overloaded (def_name, eq) lthy =
- let
- val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq;
- val ((_, (_, thm)), lthy') = lthy |> LocalTheory.define Thm.definitionK
- ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs));
- val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy');
- val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
- in (thm', lthy') end;
-
- val ((size_def_thms, size_def_thms'), thy') =
- thy
- |> Sign.add_consts_i (map (fn (s, T) =>
- (Binding.name (Long_Name.base_name s), param_size_fTs @ [T] ---> HOLogic.natT, NoSyn))
- (size_names ~~ recTs1))
- |> PureThy.add_defs false
- (map (Thm.no_attributes o apsnd (Logic.mk_equals o apsnd (app fs)))
- (map Binding.name def_names ~~ (size_fns ~~ rec_combs1)))
- ||> TheoryTarget.instantiation
- (map (#1 o snd) descr', map dest_TFree paramTs, [HOLogic.class_size])
- ||>> fold_map define_overloaded
- (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1))
- ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
- ||> LocalTheory.exit_global;
-
- val ctxt = ProofContext.init thy';
-
- val simpset1 = HOL_basic_ss addsimps @{thm add_0} :: @{thm add_0_right} ::
- size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites;
- val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2);
-
- fun mk_unfolded_size_eq tab size_ofp fs (p as (x, T), r) =
- HOLogic.mk_eq (app fs r $ Free p,
- the (size_of_type tab extra_size size_ofp T) $ Free p);
-
- fun prove_unfolded_size_eqs size_ofp fs =
- if null recTs2 then []
- else split_conj_thm (SkipProof.prove ctxt xs []
- (HOLogic.mk_Trueprop (mk_conj (replicate l HOLogic.true_const @
- map (mk_unfolded_size_eq (AList.lookup op =
- (new_type_names ~~ map (app fs) rec_combs1)) size_ofp fs)
- (xs ~~ recTs2 ~~ rec_combs2))))
- (fn _ => (indtac induction xs THEN_ALL_NEW asm_simp_tac simpset1) 1));
-
- val unfolded_size_eqs1 = prove_unfolded_size_eqs param_size fs;
- val unfolded_size_eqs2 = prove_unfolded_size_eqs (K NONE) fs';
-
- (* characteristic equations for size functions *)
- fun gen_mk_size_eq p size_of size_ofp size_const T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val tnames = Name.variant_list f_names (DatatypeProp.make_tnames Ts);
- val ts = map_filter (fn (sT as (s, T), dt) =>
- Option.map (fn sz => sz $ Free sT)
- (if p dt then size_of_type size_of extra_size size_ofp T
- else NONE)) (tnames ~~ Ts ~~ cargs)
- in
- HOLogic.mk_Trueprop (HOLogic.mk_eq
- (size_const $ list_comb (Const (cname, Ts ---> T),
- map2 (curry Free) tnames Ts),
- if null ts then HOLogic.zero
- else foldl1 plus (ts @ [HOLogic.Suc_zero])))
- end;
-
- val simpset2 = HOL_basic_ss addsimps
- rec_rewrites @ size_def_thms @ unfolded_size_eqs1;
- val simpset3 = HOL_basic_ss addsimps
- rec_rewrites @ size_def_thms' @ unfolded_size_eqs2;
-
- fun prove_size_eqs p size_fns size_ofp simpset =
- maps (fn (((_, (_, _, constrs)), size_const), T) =>
- map (fn constr => standard (SkipProof.prove ctxt [] []
- (gen_mk_size_eq p (AList.lookup op = (new_type_names ~~ size_fns))
- size_ofp size_const T constr)
- (fn _ => simp_tac simpset 1))) constrs)
- (descr' ~~ size_fns ~~ recTs1);
-
- val size_eqns = prove_size_eqs (is_poly thy') size_fns param_size simpset2 @
- prove_size_eqs is_rec_type overloaded_size_fns (K NONE) simpset3;
-
- val ([size_thms], thy'') = PureThy.add_thmss
- [((Binding.name "size", size_eqns),
- [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
- Thm.declaration_attribute
- (fn thm => Context.mapping (Code.add_default_eqn thm) I)])] thy'
-
- in
- SizeData.map (fold (Symtab.update_new o apsnd (rpair size_thms))
- (new_type_names ~~ size_names)) thy''
- end;
-
-fun add_size_thms config (new_type_names as name :: _) thy =
- let
- val info as {descr, alt_names, ...} = DatatypePackage.the_datatype thy name;
- val prefix = Long_Name.map_base_name (K (space_implode "_"
- (the_default (map Long_Name.base_name new_type_names) alt_names))) name;
- val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt =>
- is_rec_type dt andalso not (null (fst (strip_dtyp dt)))) cargs) constrs) descr
- in if no_size then thy
- else
- thy
- |> Sign.root_path
- |> Sign.add_path prefix
- |> Theory.checkpoint
- |> prove_size_thms info new_type_names
- |> Sign.restore_naming thy
- end;
-
-val size_thms = snd oo (the oo lookup_size);
-
-val setup = DatatypePackage.interpretation add_size_thms;
-
-end;
--- a/src/HOL/Tools/function_package/sum_tree.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,44 +0,0 @@
-(* Title: HOL/Tools/function_package/sum_tree.ML
- ID: $Id$
- Author: Alexander Krauss, TU Muenchen
-
-Some common tools for working with sum types in balanced tree form.
-*)
-
-structure SumTree =
-struct
-
-(* Theory dependencies *)
-val proj_in_rules = [@{thm "Datatype.Projl_Inl"}, @{thm "Datatype.Projr_Inr"}]
-val sumcase_split_ss = HOL_basic_ss addsimps (@{thm "Product_Type.split"} :: @{thms "sum.cases"})
-
-(* top-down access in balanced tree *)
-fun access_top_down {left, right, init} len i =
- BalancedTree.access {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
-
-(* Sum types *)
-fun mk_sumT LT RT = Type ("+", [LT, RT])
-fun mk_sumcase TL TR T l r = Const (@{const_name "sum.sum_case"}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
-
-val App = curry op $
-
-fun mk_inj ST n i =
- access_top_down
- { init = (ST, I : term -> term),
- left = (fn (T as Type ("+", [LT, RT]), inj) => (LT, inj o App (Const (@{const_name "Inl"}, LT --> T)))),
- right =(fn (T as Type ("+", [LT, RT]), inj) => (RT, inj o App (Const (@{const_name "Inr"}, RT --> T))))} n i
- |> snd
-
-fun mk_proj ST n i =
- access_top_down
- { init = (ST, I : term -> term),
- left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name "Datatype.Projl"}, T --> LT)) o proj)),
- right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name "Datatype.Projr"}, T --> RT)) o proj))} n i
- |> snd
-
-fun mk_sumcases T fs =
- BalancedTree.make (fn ((f, fT), (g, gT)) => (mk_sumcase fT gT T f g, mk_sumT fT gT))
- (map (fn f => (f, domain_type (fastype_of f))) fs)
- |> fst
-
-end
--- a/src/HOL/Tools/function_package/termination.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,324 +0,0 @@
-(* Title: HOL/Tools/function_package/termination.ML
- Author: Alexander Krauss, TU Muenchen
-
-Context data for termination proofs
-*)
-
-
-signature TERMINATION =
-sig
-
- type data
- datatype cell = Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm
-
- val mk_sumcases : data -> typ -> term list -> term
-
- val note_measure : int -> term -> data -> data
- val note_chain : term -> term -> thm option -> data -> data
- val note_descent : term -> term -> term -> cell -> data -> data
-
- val get_num_points : data -> int
- val get_types : data -> int -> typ
- val get_measures : data -> int -> term list
-
- (* read from cache *)
- val get_chain : data -> term -> term -> thm option option
- val get_descent : data -> term -> term -> term -> cell option
-
- (* writes *)
- val derive_descent : theory -> tactic -> term -> term -> term -> data -> data
- val derive_descents : theory -> tactic -> term -> data -> data
-
- val dest_call : data -> term -> ((string * typ) list * int * term * int * term * term)
-
- val CALLS : (term list * int -> tactic) -> int -> tactic
-
- (* Termination tactics. Sequential composition via continuations. (2nd argument is the error continuation) *)
- type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
-
- val TERMINATION : Proof.context -> (data -> int -> tactic) -> int -> tactic
-
- val REPEAT : ttac -> ttac
-
- val wf_union_tac : Proof.context -> tactic
-end
-
-
-
-structure Termination : TERMINATION =
-struct
-
-open FundefLib
-
-val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
-structure Term2tab = TableFun(type key = term * term val ord = term2_ord);
-structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord TermOrd.fast_term_ord term2_ord);
-
-(** Analyzing binary trees **)
-
-(* Skeleton of a tree structure *)
-
-datatype skel =
- SLeaf of int (* index *)
-| SBranch of (skel * skel)
-
-
-(* abstract make and dest functions *)
-fun mk_tree leaf branch =
- let fun mk (SLeaf i) = leaf i
- | mk (SBranch (s, t)) = branch (mk s, mk t)
- in mk end
-
-
-fun dest_tree split =
- let fun dest (SLeaf i) x = [(i, x)]
- | dest (SBranch (s, t)) x =
- let val (l, r) = split x
- in dest s l @ dest t r end
- in dest end
-
-
-(* concrete versions for sum types *)
-fun is_inj (Const ("Sum_Type.Inl", _) $ _) = true
- | is_inj (Const ("Sum_Type.Inr", _) $ _) = true
- | is_inj _ = false
-
-fun dest_inl (Const ("Sum_Type.Inl", _) $ t) = SOME t
- | dest_inl _ = NONE
-
-fun dest_inr (Const ("Sum_Type.Inr", _) $ t) = SOME t
- | dest_inr _ = NONE
-
-
-fun mk_skel ps =
- let
- fun skel i ps =
- if forall is_inj ps andalso not (null ps)
- then let
- val (j, s) = skel i (map_filter dest_inl ps)
- val (k, t) = skel j (map_filter dest_inr ps)
- in (k, SBranch (s, t)) end
- else (i + 1, SLeaf i)
- in
- snd (skel 0 ps)
- end
-
-(* compute list of types for nodes *)
-fun node_types sk T = dest_tree (fn Type ("+", [LT, RT]) => (LT, RT)) sk T |> map snd
-
-(* find index and raw term *)
-fun dest_inj (SLeaf i) trm = (i, trm)
- | dest_inj (SBranch (s, t)) trm =
- case dest_inl trm of
- SOME trm' => dest_inj s trm'
- | _ => dest_inj t (the (dest_inr trm))
-
-
-
-(** Matrix cell datatype **)
-
-datatype cell = Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm;
-
-
-type data =
- skel (* structure of the sum type encoding "program points" *)
- * (int -> typ) (* types of program points *)
- * (term list Inttab.table) (* measures for program points *)
- * (thm option Term2tab.table) (* which calls form chains? *)
- * (cell Term3tab.table) (* local descents *)
-
-
-fun map_measures f (p, T, M, C, D) = (p, T, f M, C, D)
-fun map_chains f (p, T, M, C, D) = (p, T, M, f C, D)
-fun map_descent f (p, T, M, C, D) = (p, T, M, C, f D)
-
-fun note_measure p m = map_measures (Inttab.insert_list (op aconv) (p, m))
-fun note_chain c1 c2 res = map_chains (Term2tab.update ((c1, c2), res))
-fun note_descent c m1 m2 res = map_descent (Term3tab.update ((c,(m1, m2)), res))
-
-(* Build case expression *)
-fun mk_sumcases (sk, _, _, _, _) T fs =
- mk_tree (fn i => (nth fs i, domain_type (fastype_of (nth fs i))))
- (fn ((f, fT), (g, gT)) => (SumTree.mk_sumcase fT gT T f g, SumTree.mk_sumT fT gT))
- sk
- |> fst
-
-fun mk_sum_skel rel =
- let
- val cs = FundefLib.dest_binop_list @{const_name Un} rel
- fun collect_pats (Const ("Collect", _) $ Abs (_, _, c)) =
- let
- val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
- = Term.strip_qnt_body "Ex" c
- in cons r o cons l end
- in
- mk_skel (fold collect_pats cs [])
- end
-
-fun create ctxt T rel =
- let
- val sk = mk_sum_skel rel
- val Ts = node_types sk T
- val M = Inttab.make (map_index (apsnd (MeasureFunctions.get_measure_functions ctxt)) Ts)
- in
- (sk, nth Ts, M, Term2tab.empty, Term3tab.empty)
- end
-
-fun get_num_points (sk, _, _, _, _) =
- let
- fun num (SLeaf i) = i + 1
- | num (SBranch (s, t)) = num t
- in num sk end
-
-fun get_types (_, T, _, _, _) = T
-fun get_measures (_, _, M, _, _) = Inttab.lookup_list M
-
-fun get_chain (_, _, _, C, _) c1 c2 =
- Term2tab.lookup C (c1, c2)
-
-fun get_descent (_, _, _, _, D) c m1 m2 =
- Term3tab.lookup D (c, (m1, m2))
-
-fun dest_call D (Const ("Collect", _) $ Abs (_, _, c)) =
- let
- val n = get_num_points D
- val (sk, _, _, _, _) = D
- val vs = Term.strip_qnt_vars "Ex" c
-
- (* FIXME: throw error "dest_call" for malformed terms *)
- val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
- = Term.strip_qnt_body "Ex" c
- val (p, l') = dest_inj sk l
- val (q, r') = dest_inj sk r
- in
- (vs, p, l', q, r', Gam)
- end
- | dest_call D t = error "dest_call"
-
-
-fun derive_desc_aux thy tac c (vs, p, l', q, r', Gam) m1 m2 D =
- case get_descent D c m1 m2 of
- SOME _ => D
- | NONE => let
- fun cgoal rel =
- Term.list_all (vs,
- Logic.mk_implies (HOLogic.mk_Trueprop Gam,
- HOLogic.mk_Trueprop (Const (rel, @{typ "nat => nat => bool"})
- $ (m2 $ r') $ (m1 $ l'))))
- |> cterm_of thy
- in
- note_descent c m1 m2
- (case try_proof (cgoal @{const_name HOL.less}) tac of
- Solved thm => Less thm
- | Stuck thm =>
- (case try_proof (cgoal @{const_name HOL.less_eq}) tac of
- Solved thm2 => LessEq (thm2, thm)
- | Stuck thm2 =>
- if prems_of thm2 = [HOLogic.Trueprop $ HOLogic.false_const]
- then False thm2 else None (thm2, thm)
- | _ => raise Match) (* FIXME *)
- | _ => raise Match) D
- end
-
-fun derive_descent thy tac c m1 m2 D =
- derive_desc_aux thy tac c (dest_call D c) m1 m2 D
-
-(* all descents in one go *)
-fun derive_descents thy tac c D =
- let val cdesc as (vs, p, l', q, r', Gam) = dest_call D c
- in fold_product (derive_desc_aux thy tac c cdesc)
- (get_measures D p) (get_measures D q) D
- end
-
-fun CALLS tac i st =
- if Thm.no_prems st then all_tac st
- else case Thm.term_of (Thm.cprem_of st i) of
- (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Un} rel, i) st
- |_ => no_tac st
-
-type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
-
-fun TERMINATION ctxt tac =
- SUBGOAL (fn (_ $ (Const (@{const_name "wf"}, wfT) $ rel), i) =>
- let
- val (T, _) = HOLogic.dest_prodT (HOLogic.dest_setT (domain_type wfT))
- in
- tac (create ctxt T rel) i
- end)
-
-
-(* A tactic to convert open to closed termination goals *)
-local
-fun dest_term (t : term) = (* FIXME, cf. Lexicographic order *)
- let
- val (vars, prop) = FundefLib.dest_all_all t
- val (prems, concl) = Logic.strip_horn prop
- val (lhs, rhs) = concl
- |> HOLogic.dest_Trueprop
- |> HOLogic.dest_mem |> fst
- |> HOLogic.dest_prod
- in
- (vars, prems, lhs, rhs)
- end
-
-fun mk_pair_compr (T, qs, l, r, conds) =
- let
- val pT = HOLogic.mk_prodT (T, T)
- val n = length qs
- val peq = HOLogic.eq_const pT $ Bound n $ (HOLogic.pair_const T T $ l $ r)
- val conds' = if null conds then [HOLogic.true_const] else conds
- in
- HOLogic.Collect_const pT $
- Abs ("uu_", pT,
- (foldr1 HOLogic.mk_conj (peq :: conds')
- |> fold_rev (fn v => fn t => HOLogic.exists_const (fastype_of v) $ lambda v t) qs))
- end
-
-in
-
-fun wf_union_tac ctxt st =
- let
- val thy = ProofContext.theory_of ctxt
- val cert = cterm_of (theory_of_thm st)
- val ((trueprop $ (wf $ rel)) :: ineqs) = prems_of st
-
- fun mk_compr ineq =
- let
- val (vars, prems, lhs, rhs) = dest_term ineq
- in
- mk_pair_compr (fastype_of lhs, vars, lhs, rhs, map (ObjectLogic.atomize_term thy) prems)
- end
-
- val relation =
- if null ineqs then
- Const (@{const_name Set.empty}, fastype_of rel)
- else
- foldr1 (HOLogic.mk_binop @{const_name Un}) (map mk_compr ineqs)
-
- fun solve_membership_tac i =
- (EVERY' (replicate (i - 2) (rtac @{thm UnI2})) (* pick the right component of the union *)
- THEN' (fn j => TRY (rtac @{thm UnI1} j))
- THEN' (rtac @{thm CollectI}) (* unfold comprehension *)
- THEN' (fn i => REPEAT (rtac @{thm exI} i)) (* Turn existentials into schematic Vars *)
- THEN' ((rtac @{thm refl}) (* unification instantiates all Vars *)
- ORELSE' ((rtac @{thm conjI})
- THEN' (rtac @{thm refl})
- THEN' (blast_tac (local_claset_of ctxt)))) (* Solve rest of context... not very elegant *)
- ) i
- in
- ((PRIMITIVE (Drule.cterm_instantiate [(cert rel, cert relation)])
- THEN ALLGOALS (fn i => if i = 1 then all_tac else solve_membership_tac i))) st
- end
-
-
-end
-
-
-(* continuation passing repeat combinator *)
-fun REPEAT ttac cont err_cont =
- ttac (fn D => fn i => (REPEAT ttac cont cont D i)) err_cont
-
-
-
-
-end
--- a/src/HOL/Tools/hologic.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/hologic.ML Tue Jun 23 15:48:55 2009 +0100
@@ -603,8 +603,11 @@
val typerepT = Type ("Typerep.typerep", []);
-fun mk_typerep T = Const ("Typerep.typerep_class.typerep",
- Term.itselfT T --> typerepT) $ Logic.mk_type T;
+fun mk_typerep (Type (tyco, Ts)) = Const ("Typerep.typerep.Typerep",
+ literalT --> listT typerepT --> typerepT) $ mk_literal tyco
+ $ mk_list typerepT (map mk_typerep Ts)
+ | mk_typerep (T as TFree _) = Const ("Typerep.typerep_class.typerep",
+ Term.itselfT T --> typerepT) $ Logic.mk_type T;
val termT = Type ("Code_Eval.term", []);
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/inductive.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,968 @@
+(* Title: HOL/Tools/inductive.ML
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
+
+(Co)Inductive Definition module for HOL.
+
+Features:
+ * least or greatest fixedpoints
+ * mutually recursive definitions
+ * definitions involving arbitrary monotone operators
+ * automatically proves introduction and elimination rules
+
+ Introduction rules have the form
+ [| M Pj ti, ..., Q x, ... |] ==> Pk t
+ where M is some monotone operator (usually the identity)
+ Q x is any side condition on the free variables
+ ti, t are any terms
+ Pj, Pk are two of the predicates being defined in mutual recursion
+*)
+
+signature BASIC_INDUCTIVE =
+sig
+ type inductive_result
+ val morph_result: morphism -> inductive_result -> inductive_result
+ type inductive_info
+ val the_inductive: Proof.context -> string -> inductive_info
+ val print_inductives: Proof.context -> unit
+ val mono_add: attribute
+ val mono_del: attribute
+ val get_monos: Proof.context -> thm list
+ val mk_cases: Proof.context -> term -> thm
+ val inductive_forall_name: string
+ val inductive_forall_def: thm
+ val rulify: thm -> thm
+ val inductive_cases: (Attrib.binding * string list) list -> local_theory ->
+ thm list list * local_theory
+ val inductive_cases_i: (Attrib.binding * term list) list -> local_theory ->
+ thm list list * local_theory
+ type inductive_flags
+ val add_inductive_i:
+ inductive_flags -> ((binding * typ) * mixfix) list ->
+ (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
+ inductive_result * local_theory
+ val add_inductive: bool -> bool ->
+ (binding * string option * mixfix) list ->
+ (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list ->
+ (Facts.ref * Attrib.src list) list ->
+ bool -> local_theory -> inductive_result * local_theory
+ val add_inductive_global: string -> inductive_flags ->
+ ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+ thm list -> theory -> inductive_result * theory
+ val arities_of: thm -> (string * int) list
+ val params_of: thm -> term list
+ val partition_rules: thm -> thm list -> (string * thm list) list
+ val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list
+ val unpartition_rules: thm list -> (string * 'a list) list -> 'a list
+ val infer_intro_vars: thm -> int -> thm list -> term list list
+ val setup: theory -> theory
+end;
+
+signature INDUCTIVE =
+sig
+ include BASIC_INDUCTIVE
+ type add_ind_def
+ val declare_rules: string -> binding -> bool -> bool -> string list ->
+ thm list -> binding list -> Attrib.src list list -> (thm * string list) list ->
+ thm -> local_theory -> thm list * thm list * thm * local_theory
+ val add_ind_def: add_ind_def
+ val gen_add_inductive_i: add_ind_def -> inductive_flags ->
+ ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+ thm list -> local_theory -> inductive_result * local_theory
+ val gen_add_inductive: add_ind_def -> bool -> bool ->
+ (binding * string option * mixfix) list ->
+ (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
+ bool -> local_theory -> inductive_result * local_theory
+ val gen_ind_decl: add_ind_def -> bool ->
+ OuterParse.token list -> (bool -> local_theory -> local_theory) * OuterParse.token list
+end;
+
+structure Inductive: INDUCTIVE =
+struct
+
+
+(** theory context references **)
+
+val inductive_forall_name = "HOL.induct_forall";
+val inductive_forall_def = thm "induct_forall_def";
+val inductive_conj_name = "HOL.induct_conj";
+val inductive_conj_def = thm "induct_conj_def";
+val inductive_conj = thms "induct_conj";
+val inductive_atomize = thms "induct_atomize";
+val inductive_rulify = thms "induct_rulify";
+val inductive_rulify_fallback = thms "induct_rulify_fallback";
+
+val notTrueE = TrueI RSN (2, notE);
+val notFalseI = Seq.hd (atac 1 notI);
+val simp_thms' = map (fn s => mk_meta_eq (the (find_first
+ (equal (OldGoals.read_prop @{theory HOL} s) o prop_of) simp_thms)))
+ ["(~True) = False", "(~False) = True",
+ "(True --> ?P) = ?P", "(False --> ?P) = True",
+ "(?P & True) = ?P", "(True & ?P) = ?P"];
+
+
+
+(** context data **)
+
+type inductive_result =
+ {preds: term list, elims: thm list, raw_induct: thm,
+ induct: thm, intrs: thm list};
+
+fun morph_result phi {preds, elims, raw_induct: thm, induct, intrs} =
+ let
+ val term = Morphism.term phi;
+ val thm = Morphism.thm phi;
+ val fact = Morphism.fact phi;
+ in
+ {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct,
+ induct = thm induct, intrs = fact intrs}
+ end;
+
+type inductive_info =
+ {names: string list, coind: bool} * inductive_result;
+
+structure InductiveData = GenericDataFun
+(
+ type T = inductive_info Symtab.table * thm list;
+ val empty = (Symtab.empty, []);
+ val extend = I;
+ fun merge _ ((tab1, monos1), (tab2, monos2)) =
+ (Symtab.merge (K true) (tab1, tab2), Thm.merge_thms (monos1, monos2));
+);
+
+val get_inductives = InductiveData.get o Context.Proof;
+
+fun print_inductives ctxt =
+ let
+ val (tab, monos) = get_inductives ctxt;
+ val space = Consts.space_of (ProofContext.consts_of ctxt);
+ in
+ [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))),
+ Pretty.big_list "monotonicity rules:" (map (ProofContext.pretty_thm ctxt) monos)]
+ |> Pretty.chunks |> Pretty.writeln
+ end;
+
+
+(* get and put data *)
+
+fun the_inductive ctxt name =
+ (case Symtab.lookup (#1 (get_inductives ctxt)) name of
+ NONE => error ("Unknown (co)inductive predicate " ^ quote name)
+ | SOME info => info);
+
+fun put_inductives names info = InductiveData.map
+ (apfst (fold (fn name => Symtab.update (name, info)) names));
+
+
+
+(** monotonicity rules **)
+
+val get_monos = #2 o get_inductives;
+val map_monos = InductiveData.map o apsnd;
+
+fun mk_mono thm =
+ let
+ val concl = concl_of thm;
+ fun eq2mono thm' = [thm' RS (thm' RS eq_to_mono)] @
+ (case concl of
+ (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
+ | _ => [thm' RS (thm' RS eq_to_mono2)]);
+ fun dest_less_concl thm = dest_less_concl (thm RS le_funD)
+ handle THM _ => thm RS le_boolD
+ in
+ case concl of
+ Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
+ | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
+ | _ $ (Const ("HOL.ord_class.less_eq", _) $ _ $ _) =>
+ [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
+ (resolve_tac [le_funI, le_boolI'])) thm))]
+ | _ => [thm]
+ end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm thm);
+
+val mono_add = Thm.declaration_attribute (map_monos o fold Thm.add_thm o mk_mono);
+val mono_del = Thm.declaration_attribute (map_monos o fold Thm.del_thm o mk_mono);
+
+
+
+(** misc utilities **)
+
+fun message quiet_mode s = if quiet_mode then () else writeln s;
+fun clean_message quiet_mode s = if ! quick_and_dirty then () else message quiet_mode s;
+
+fun coind_prefix true = "co"
+ | coind_prefix false = "";
+
+fun log (b:int) m n = if m >= n then 0 else 1 + log b (b * m) n;
+
+fun make_bool_args f g [] i = []
+ | make_bool_args f g (x :: xs) i =
+ (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2);
+
+fun make_bool_args' xs =
+ make_bool_args (K HOLogic.false_const) (K HOLogic.true_const) xs;
+
+fun find_arg T x [] = sys_error "find_arg"
+ | find_arg T x ((p as (_, (SOME _, _))) :: ps) =
+ apsnd (cons p) (find_arg T x ps)
+ | find_arg T x ((p as (U, (NONE, y))) :: ps) =
+ if (T: typ) = U then (y, (U, (SOME x, y)) :: ps)
+ else apsnd (cons p) (find_arg T x ps);
+
+fun make_args Ts xs =
+ map (fn (T, (NONE, ())) => Const (@{const_name undefined}, T) | (_, (SOME t, ())) => t)
+ (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts));
+
+fun make_args' Ts xs Us =
+ fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs));
+
+fun dest_predicate cs params t =
+ let
+ val k = length params;
+ val (c, ts) = strip_comb t;
+ val (xs, ys) = chop k ts;
+ val i = find_index_eq c cs;
+ in
+ if xs = params andalso i >= 0 then
+ SOME (c, i, ys, chop (length ys)
+ (List.drop (binder_types (fastype_of c), k)))
+ else NONE
+ end;
+
+fun mk_names a 0 = []
+ | mk_names a 1 = [a]
+ | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
+
+
+
+(** process rules **)
+
+local
+
+fun err_in_rule ctxt name t msg =
+ error (cat_lines ["Ill-formed introduction rule " ^ quote name,
+ Syntax.string_of_term ctxt t, msg]);
+
+fun err_in_prem ctxt name t p msg =
+ error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p,
+ "in introduction rule " ^ quote name, Syntax.string_of_term ctxt t, msg]);
+
+val bad_concl = "Conclusion of introduction rule must be an inductive predicate";
+
+val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate";
+
+val bad_app = "Inductive predicate must be applied to parameter(s) ";
+
+fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
+
+in
+
+fun check_rule ctxt cs params ((binding, att), rule) =
+ let
+ val err_name = Binding.str_of binding;
+ val params' = Term.variant_frees rule (Logic.strip_params rule);
+ val frees = rev (map Free params');
+ val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
+ val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule);
+ val rule' = Logic.list_implies (prems, concl);
+ val aprems = map (atomize_term (ProofContext.theory_of ctxt)) prems;
+ val arule = list_all_free (params', Logic.list_implies (aprems, concl));
+
+ fun check_ind err t = case dest_predicate cs params t of
+ NONE => err (bad_app ^
+ commas (map (Syntax.string_of_term ctxt) params))
+ | SOME (_, _, ys, _) =>
+ if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs
+ then err bad_ind_occ else ();
+
+ fun check_prem' prem t =
+ if head_of t mem cs then
+ check_ind (err_in_prem ctxt err_name rule prem) t
+ else (case t of
+ Abs (_, _, t) => check_prem' prem t
+ | t $ u => (check_prem' prem t; check_prem' prem u)
+ | _ => ());
+
+ fun check_prem (prem, aprem) =
+ if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
+ else err_in_prem ctxt err_name rule prem "Non-atomic premise";
+ in
+ (case concl of
+ Const ("Trueprop", _) $ t =>
+ if head_of t mem cs then
+ (check_ind (err_in_rule ctxt err_name rule') t;
+ List.app check_prem (prems ~~ aprems))
+ else err_in_rule ctxt err_name rule' bad_concl
+ | _ => err_in_rule ctxt err_name rule' bad_concl);
+ ((binding, att), arule)
+ end;
+
+val rulify =
+ hol_simplify inductive_conj
+ #> hol_simplify inductive_rulify
+ #> hol_simplify inductive_rulify_fallback
+ #> Simplifier.norm_hhf;
+
+end;
+
+
+
+(** proofs for (co)inductive predicates **)
+
+(* prove monotonicity *)
+
+fun prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt =
+ (message (quiet_mode orelse skip_mono andalso !quick_and_dirty orelse fork_mono)
+ " Proving monotonicity ...";
+ (if skip_mono then SkipProof.prove else if fork_mono then Goal.prove_future else Goal.prove) ctxt
+ [] []
+ (HOLogic.mk_Trueprop
+ (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
+ (fn _ => EVERY [rtac @{thm monoI} 1,
+ REPEAT (resolve_tac [le_funI, le_boolI'] 1),
+ REPEAT (FIRST
+ [atac 1,
+ resolve_tac (List.concat (map mk_mono monos) @ get_monos ctxt) 1,
+ etac le_funE 1, dtac le_boolD 1])]));
+
+
+(* prove introduction rules *)
+
+fun prove_intrs quiet_mode coind mono fp_def k params intr_ts rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the introduction rules ...";
+
+ val unfold = funpow k (fn th => th RS fun_cong)
+ (mono RS (fp_def RS
+ (if coind then def_gfp_unfold else def_lfp_unfold)));
+
+ fun select_disj 1 1 = []
+ | select_disj _ 1 = [rtac disjI1]
+ | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
+
+ val rules = [refl, TrueI, notFalseI, exI, conjI];
+
+ val intrs = map_index (fn (i, intr) => rulify
+ (SkipProof.prove ctxt (map (fst o dest_Free) params) [] intr (fn _ => EVERY
+ [rewrite_goals_tac rec_preds_defs,
+ rtac (unfold RS iffD2) 1,
+ EVERY1 (select_disj (length intr_ts) (i + 1)),
+ (*Not ares_tac, since refl must be tried before any equality assumptions;
+ backtracking may occur if the premises have extra variables!*)
+ DEPTH_SOLVE_1 (resolve_tac rules 1 APPEND assume_tac 1)]))) intr_ts
+
+ in (intrs, unfold) end;
+
+
+(* prove elimination rules *)
+
+fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the elimination rules ...";
+
+ val ([pname], ctxt') = ctxt |>
+ Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
+ Variable.variant_fixes ["P"];
+ val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
+
+ fun dest_intr r =
+ (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
+ Logic.strip_assums_hyp r, Logic.strip_params r);
+
+ val intrs = map dest_intr intr_ts ~~ intr_names;
+
+ val rules1 = [disjE, exE, FalseE];
+ val rules2 = [conjE, FalseE, notTrueE];
+
+ fun prove_elim c =
+ let
+ val Ts = List.drop (binder_types (fastype_of c), length params);
+ val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt';
+ val frees = map Free (anames ~~ Ts);
+
+ fun mk_elim_prem ((_, _, us, _), ts, params') =
+ list_all (params',
+ Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
+ (frees ~~ us) @ ts, P));
+ val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs);
+ val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
+ map mk_elim_prem (map #1 c_intrs)
+ in
+ (SkipProof.prove ctxt'' [] prems P
+ (fn {prems, ...} => EVERY
+ [cut_facts_tac [hd prems] 1,
+ rewrite_goals_tac rec_preds_defs,
+ dtac (unfold RS iffD1) 1,
+ REPEAT (FIRSTGOAL (eresolve_tac rules1)),
+ REPEAT (FIRSTGOAL (eresolve_tac rules2)),
+ EVERY (map (fn prem =>
+ DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_preds_defs prem, conjI] 1)) (tl prems))])
+ |> rulify
+ |> singleton (ProofContext.export ctxt'' ctxt),
+ map #2 c_intrs)
+ end
+
+ in map prove_elim cs end;
+
+
+(* derivation of simplified elimination rules *)
+
+local
+
+(*delete needless equality assumptions*)
+val refl_thin = Goal.prove_global @{theory HOL} [] [] @{prop "!!P. a = a ==> P ==> P"}
+ (fn _ => assume_tac 1);
+val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
+val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
+
+fun simp_case_tac ss i =
+ EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i;
+
+in
+
+fun mk_cases ctxt prop =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val ss = Simplifier.local_simpset_of ctxt;
+
+ fun err msg =
+ error (Pretty.string_of (Pretty.block
+ [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop]));
+
+ val elims = Induct.find_casesP ctxt prop;
+
+ val cprop = Thm.cterm_of thy prop;
+ val tac = ALLGOALS (simp_case_tac ss) THEN prune_params_tac;
+ fun mk_elim rl =
+ Thm.implies_intr cprop (Tactic.rule_by_tactic tac (Thm.assume cprop RS rl))
+ |> singleton (Variable.export (Variable.auto_fixes prop ctxt) ctxt);
+ in
+ (case get_first (try mk_elim) elims of
+ SOME r => r
+ | NONE => err "Proposition not an inductive predicate:")
+ end;
+
+end;
+
+
+(* inductive_cases *)
+
+fun gen_inductive_cases prep_att prep_prop args lthy =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val facts = args |> map (fn ((a, atts), props) =>
+ ((a, map (prep_att thy) atts),
+ map (Thm.no_attributes o single o mk_cases lthy o prep_prop lthy) props));
+ in lthy |> LocalTheory.notes Thm.generatedK facts |>> map snd end;
+
+val inductive_cases = gen_inductive_cases Attrib.intern_src Syntax.read_prop;
+val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
+
+
+val ind_cases_setup =
+ Method.setup @{binding ind_cases}
+ (Scan.lift (Scan.repeat1 Args.name_source --
+ Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.name) []) >>
+ (fn (raw_props, fixes) => fn ctxt =>
+ let
+ val (_, ctxt') = Variable.add_fixes fixes ctxt;
+ val props = Syntax.read_props ctxt' raw_props;
+ val ctxt'' = fold Variable.declare_term props ctxt';
+ val rules = ProofContext.export ctxt'' ctxt (map (mk_cases ctxt'') props)
+ in Method.erule 0 rules end))
+ "dynamic case analysis on predicates";
+
+
+(* prove induction rule *)
+
+fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono
+ fp_def rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the induction rule ...";
+ val thy = ProofContext.theory_of ctxt;
+
+ (* predicates for induction rule *)
+
+ val (pnames, ctxt') = ctxt |>
+ Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
+ Variable.variant_fixes (mk_names "P" (length cs));
+ val preds = map Free (pnames ~~
+ map (fn c => List.drop (binder_types (fastype_of c), length params) --->
+ HOLogic.boolT) cs);
+
+ (* transform an introduction rule into a premise for induction rule *)
+
+ fun mk_ind_prem r =
+ let
+ fun subst s = (case dest_predicate cs params s of
+ SOME (_, i, ys, (_, Ts)) =>
+ let
+ val k = length Ts;
+ val bs = map Bound (k - 1 downto 0);
+ val P = list_comb (List.nth (preds, i),
+ map (incr_boundvars k) ys @ bs);
+ val Q = list_abs (mk_names "x" k ~~ Ts,
+ HOLogic.mk_binop inductive_conj_name
+ (list_comb (incr_boundvars k s, bs), P))
+ in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end
+ | NONE => (case s of
+ (t $ u) => (fst (subst t) $ fst (subst u), NONE)
+ | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
+ | _ => (s, NONE)));
+
+ fun mk_prem (s, prems) = (case subst s of
+ (_, SOME (t, u)) => t :: u :: prems
+ | (t, _) => t :: prems);
+
+ val SOME (_, i, ys, _) = dest_predicate cs params
+ (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))
+
+ in list_all_free (Logic.strip_params r,
+ Logic.list_implies (map HOLogic.mk_Trueprop (List.foldr mk_prem
+ [] (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r))),
+ HOLogic.mk_Trueprop (list_comb (List.nth (preds, i), ys))))
+ end;
+
+ val ind_prems = map mk_ind_prem intr_ts;
+
+
+ (* make conclusions for induction rules *)
+
+ val Tss = map (binder_types o fastype_of) preds;
+ val (xnames, ctxt'') =
+ Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt';
+ val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((xnames, Ts), c), P) =>
+ let val frees = map Free (xnames ~~ Ts)
+ in HOLogic.mk_imp
+ (list_comb (c, params @ frees), list_comb (P, frees))
+ end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds)));
+
+
+ (* make predicate for instantiation of abstract induction rule *)
+
+ val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj
+ (map_index (fn (i, P) => List.foldr HOLogic.mk_imp
+ (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))
+ (make_bool_args HOLogic.mk_not I bs i)) preds));
+
+ val ind_concl = HOLogic.mk_Trueprop
+ (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred));
+
+ val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
+
+ val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
+ (fn {prems, ...} => EVERY
+ [rewrite_goals_tac [inductive_conj_def],
+ DETERM (rtac raw_fp_induct 1),
+ REPEAT (resolve_tac [le_funI, le_boolI] 1),
+ rewrite_goals_tac (inf_fun_eq :: inf_bool_eq :: simp_thms'),
+ (*This disjE separates out the introduction rules*)
+ REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
+ (*Now break down the individual cases. No disjE here in case
+ some premise involves disjunction.*)
+ REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
+ REPEAT (FIRSTGOAL
+ (resolve_tac [conjI, impI] ORELSE' (etac notE THEN' atac))),
+ EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule
+ (inductive_conj_def :: rec_preds_defs @ simp_thms') prem,
+ conjI, refl] 1)) prems)]);
+
+ val lemma = SkipProof.prove ctxt'' [] []
+ (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
+ [rewrite_goals_tac rec_preds_defs,
+ REPEAT (EVERY
+ [REPEAT (resolve_tac [conjI, impI] 1),
+ REPEAT (eresolve_tac [le_funE, le_boolE] 1),
+ atac 1,
+ rewrite_goals_tac simp_thms',
+ atac 1])])
+
+ in singleton (ProofContext.export ctxt'' ctxt) (induct RS lemma) end;
+
+
+
+(** specification of (co)inductive predicates **)
+
+fun mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts monos params cnames_syn ctxt =
+ let
+ val fp_name = if coind then @{const_name Inductive.gfp} else @{const_name Inductive.lfp};
+
+ val argTs = fold (fn c => fn Ts => Ts @
+ (List.drop (binder_types (fastype_of c), length params) \\ Ts)) cs [];
+ val k = log 2 1 (length cs);
+ val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT;
+ val p :: xs = map Free (Variable.variant_frees ctxt intr_ts
+ (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs)));
+ val bs = map Free (Variable.variant_frees ctxt (p :: xs @ intr_ts)
+ (map (rpair HOLogic.boolT) (mk_names "b" k)));
+
+ fun subst t = (case dest_predicate cs params t of
+ SOME (_, i, ts, (Ts, Us)) =>
+ let
+ val l = length Us;
+ val zs = map Bound (l - 1 downto 0)
+ in
+ list_abs (map (pair "z") Us, list_comb (p,
+ make_bool_args' bs i @ make_args argTs
+ ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us))))
+ end
+ | NONE => (case t of
+ t1 $ t2 => subst t1 $ subst t2
+ | Abs (x, T, u) => Abs (x, T, subst u)
+ | _ => t));
+
+ (* transform an introduction rule into a conjunction *)
+ (* [| p_i t; ... |] ==> p_j u *)
+ (* is transformed into *)
+ (* b_j & x_j = u & p b_j t & ... *)
+
+ fun transform_rule r =
+ let
+ val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params
+ (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
+ val ps = make_bool_args HOLogic.mk_not I bs i @
+ map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
+ map (subst o HOLogic.dest_Trueprop)
+ (Logic.strip_assums_hyp r)
+ in List.foldr (fn ((x, T), P) => HOLogic.exists_const T $ (Abs (x, T, P)))
+ (if null ps then HOLogic.true_const else foldr1 HOLogic.mk_conj ps)
+ (Logic.strip_params r)
+ end
+
+ (* make a disjunction of all introduction rules *)
+
+ val fp_fun = fold_rev lambda (p :: bs @ xs)
+ (if null intr_ts then HOLogic.false_const
+ else foldr1 HOLogic.mk_disj (map transform_rule intr_ts));
+
+ (* add definiton of recursive predicates to theory *)
+
+ val rec_name =
+ if Binding.is_empty alt_name then
+ Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
+ else alt_name;
+
+ val ((rec_const, (_, fp_def)), ctxt') = ctxt |>
+ LocalTheory.define Thm.internalK
+ ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn),
+ (Attrib.empty_binding, fold_rev lambda params
+ (Const (fp_name, (predT --> predT) --> predT) $ fp_fun)));
+ val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def])
+ (cterm_of (ProofContext.theory_of ctxt') (list_comb (rec_const, params)));
+ val specs = if length cs < 2 then [] else
+ map_index (fn (i, (name_mx, c)) =>
+ let
+ val Ts = List.drop (binder_types (fastype_of c), length params);
+ val xs = map Free (Variable.variant_frees ctxt intr_ts
+ (mk_names "x" (length Ts) ~~ Ts))
+ in
+ (name_mx, (Attrib.empty_binding, fold_rev lambda (params @ xs)
+ (list_comb (rec_const, params @ make_bool_args' bs i @
+ make_args argTs (xs ~~ Ts)))))
+ end) (cnames_syn ~~ cs);
+ val (consts_defs, ctxt'') = fold_map (LocalTheory.define Thm.internalK) specs ctxt';
+ val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
+
+ val mono = prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt'';
+ val ((_, [mono']), ctxt''') =
+ LocalTheory.note Thm.internalK (Attrib.empty_binding, [mono]) ctxt'';
+
+ in (ctxt''', rec_name, mono', fp_def', map (#2 o #2) consts_defs,
+ list_comb (rec_const, params), preds, argTs, bs, xs)
+ end;
+
+fun declare_rules kind rec_binding coind no_ind cnames intrs intr_bindings intr_atts
+ elims raw_induct ctxt =
+ let
+ val rec_name = Binding.name_of rec_binding;
+ val rec_qualified = Binding.qualify false rec_name;
+ val intr_names = map Binding.name_of intr_bindings;
+ val ind_case_names = RuleCases.case_names intr_names;
+ val induct =
+ if coind then
+ (raw_induct, [RuleCases.case_names [rec_name],
+ RuleCases.case_conclusion (rec_name, intr_names),
+ RuleCases.consumes 1, Induct.coinduct_pred (hd cnames)])
+ else if no_ind orelse length cnames > 1 then
+ (raw_induct, [ind_case_names, RuleCases.consumes 0])
+ else (raw_induct RSN (2, rev_mp), [ind_case_names, RuleCases.consumes 1]);
+
+ val (intrs', ctxt1) =
+ ctxt |>
+ LocalTheory.notes kind
+ (map rec_qualified intr_bindings ~~ intr_atts ~~ map (fn th => [([th],
+ [Attrib.internal (K (ContextRules.intro_query NONE)),
+ Attrib.internal (K Nitpick_Ind_Intro_Thms.add)])]) intrs) |>>
+ map (hd o snd);
+ val (((_, elims'), (_, [induct'])), ctxt2) =
+ ctxt1 |>
+ LocalTheory.note kind ((rec_qualified (Binding.name "intros"), []), intrs') ||>>
+ fold_map (fn (name, (elim, cases)) =>
+ LocalTheory.note kind ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "cases"),
+ [Attrib.internal (K (RuleCases.case_names cases)),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.cases_pred name)),
+ Attrib.internal (K (ContextRules.elim_query NONE))]), [elim]) #>
+ apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
+ LocalTheory.note kind
+ ((rec_qualified (Binding.name (coind_prefix coind ^ "induct")),
+ map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]);
+
+ val ctxt3 = if no_ind orelse coind then ctxt2 else
+ let val inducts = cnames ~~ ProjectRule.projects ctxt2 (1 upto length cnames) induct'
+ in
+ ctxt2 |>
+ LocalTheory.notes kind [((rec_qualified (Binding.name "inducts"), []),
+ inducts |> map (fn (name, th) => ([th],
+ [Attrib.internal (K ind_case_names),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.induct_pred name))])))] |> snd
+ end
+ in (intrs', elims', induct', ctxt3) end;
+
+type inductive_flags =
+ {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding,
+ coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}
+
+type add_ind_def =
+ inductive_flags ->
+ term list -> (Attrib.binding * term) list -> thm list ->
+ term list -> (binding * mixfix) list ->
+ local_theory -> inductive_result * local_theory
+
+fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
+ cs intros monos params cnames_syn ctxt =
+ let
+ val _ = null cnames_syn andalso error "No inductive predicates given";
+ val names = map (Binding.name_of o fst) cnames_syn;
+ val _ = message (quiet_mode andalso not verbose)
+ ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names);
+
+ val cnames = map (LocalTheory.full_name ctxt o #1) cnames_syn; (* FIXME *)
+ val ((intr_names, intr_atts), intr_ts) =
+ apfst split_list (split_list (map (check_rule ctxt cs params) intros));
+
+ val (ctxt1, rec_name, mono, fp_def, rec_preds_defs, rec_const, preds,
+ argTs, bs, xs) = mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts
+ monos params cnames_syn ctxt;
+
+ val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
+ params intr_ts rec_preds_defs ctxt1;
+ val elims = if no_elim then [] else
+ prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names)
+ unfold rec_preds_defs ctxt1;
+ val raw_induct = zero_var_indexes
+ (if no_ind then Drule.asm_rl else
+ if coind then
+ singleton (ProofContext.export
+ (snd (Variable.add_fixes (map (fst o dest_Free) params) ctxt1)) ctxt1)
+ (rotate_prems ~1 (ObjectLogic.rulify
+ (fold_rule rec_preds_defs
+ (rewrite_rule [le_fun_def, le_bool_def, sup_fun_eq, sup_bool_eq]
+ (mono RS (fp_def RS def_coinduct))))))
+ else
+ prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
+ rec_preds_defs ctxt1);
+
+ val (intrs', elims', induct, ctxt2) = declare_rules kind rec_name coind no_ind
+ cnames intrs intr_names intr_atts elims raw_induct ctxt1;
+
+ val result =
+ {preds = preds,
+ intrs = intrs',
+ elims = elims',
+ raw_induct = rulify raw_induct,
+ induct = induct};
+
+ val ctxt3 = ctxt2
+ |> LocalTheory.declaration (fn phi =>
+ let val result' = morph_result phi result;
+ in put_inductives cnames (*global names!?*) ({names = cnames, coind = coind}, result') end);
+ in (result, ctxt3) end;
+
+
+(* external interfaces *)
+
+fun gen_add_inductive_i mk_def
+ (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono})
+ cnames_syn pnames spec monos lthy =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
+
+
+ (* abbrevs *)
+
+ val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy;
+
+ fun get_abbrev ((name, atts), t) =
+ if can (Logic.strip_assums_concl #> Logic.dest_equals) t then
+ let
+ val _ = Binding.is_empty name andalso null atts orelse
+ error "Abbreviations may not have names or attributes";
+ val ((x, T), rhs) = LocalDefs.abs_def (snd (LocalDefs.cert_def ctxt1 t));
+ val var =
+ (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of
+ NONE => error ("Undeclared head of abbreviation " ^ quote x)
+ | SOME ((b, T'), mx) =>
+ if T <> T' then error ("Bad type specification for abbreviation " ^ quote x)
+ else (b, mx));
+ in SOME (var, rhs) end
+ else NONE;
+
+ val abbrevs = map_filter get_abbrev spec;
+ val bs = map (Binding.name_of o fst o fst) abbrevs;
+
+
+ (* predicates *)
+
+ val pre_intros = filter_out (is_some o get_abbrev) spec;
+ val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn;
+ val cs = map (Free o apfst Binding.name_of o fst) cnames_syn';
+ val ps = map Free pnames;
+
+ val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn');
+ val _ = map (fn abbr => LocalDefs.fixed_abbrev abbr ctxt2) abbrevs;
+ val ctxt3 = ctxt2 |> fold (snd oo LocalDefs.fixed_abbrev) abbrevs;
+ val expand = Assumption.export_term ctxt3 lthy #> ProofContext.cert_term lthy;
+
+ fun close_rule r = list_all_free (rev (fold_aterms
+ (fn t as Free (v as (s, _)) =>
+ if Variable.is_fixed ctxt1 s orelse
+ member (op =) ps t then I else insert (op =) v
+ | _ => I) r []), r);
+
+ val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros;
+ val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn';
+ in
+ lthy
+ |> mk_def flags cs intros monos ps preds
+ ||> fold (snd oo LocalTheory.abbrev Syntax.mode_default) abbrevs
+ end;
+
+fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy =
+ let
+ val ((vars, intrs), _) = lthy
+ |> ProofContext.set_mode ProofContext.mode_abbrev
+ |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs;
+ val (cs, ps) = chop (length cnames_syn) vars;
+ val monos = Attrib.eval_thms lthy raw_monos;
+ val flags = {quiet_mode = false, verbose = verbose, kind = Thm.generatedK,
+ alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
+ skip_mono = false, fork_mono = not int};
+ in
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
+ end;
+
+val add_inductive_i = gen_add_inductive_i add_ind_def;
+val add_inductive = gen_add_inductive add_ind_def;
+
+fun add_inductive_global group flags cnames_syn pnames pre_intros monos thy =
+ let
+ val name = Sign.full_name thy (fst (fst (hd cnames_syn)));
+ val ctxt' = thy
+ |> TheoryTarget.init NONE
+ |> LocalTheory.set_group group
+ |> add_inductive_i flags cnames_syn pnames pre_intros monos |> snd
+ |> LocalTheory.exit;
+ val info = #2 (the_inductive ctxt' name);
+ in (info, ProofContext.theory_of ctxt') end;
+
+
+(* read off arities of inductive predicates from raw induction rule *)
+fun arities_of induct =
+ map (fn (_ $ t $ u) =>
+ (fst (dest_Const (head_of t)), length (snd (strip_comb u))))
+ (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+(* read off parameters of inductive predicate from raw induction rule *)
+fun params_of induct =
+ let
+ val (_ $ t $ u :: _) =
+ HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct));
+ val (_, ts) = strip_comb t;
+ val (_, us) = strip_comb u
+ in
+ List.take (ts, length ts - length us)
+ end;
+
+val pname_of_intr =
+ concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst;
+
+(* partition introduction rules according to predicate name *)
+fun gen_partition_rules f induct intros =
+ fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros
+ (map (rpair [] o fst) (arities_of induct));
+
+val partition_rules = gen_partition_rules I;
+fun partition_rules' induct = gen_partition_rules fst induct;
+
+fun unpartition_rules intros xs =
+ fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r)
+ (fn x :: xs => (x, xs)) #>> the) intros xs |> fst;
+
+(* infer order of variables in intro rules from order of quantifiers in elim rule *)
+fun infer_intro_vars elim arity intros =
+ let
+ val thy = theory_of_thm elim;
+ val _ :: cases = prems_of elim;
+ val used = map (fst o fst) (Term.add_vars (prop_of elim) []);
+ fun mtch (t, u) =
+ let
+ val params = Logic.strip_params t;
+ val vars = map (Var o apfst (rpair 0))
+ (Name.variant_list used (map fst params) ~~ map snd params);
+ val ts = map (curry subst_bounds (rev vars))
+ (List.drop (Logic.strip_assums_hyp t, arity));
+ val us = Logic.strip_imp_prems u;
+ val tab = fold (Pattern.first_order_match thy) (ts ~~ us)
+ (Vartab.empty, Vartab.empty);
+ in
+ map (Envir.subst_vars tab) vars
+ end
+ in
+ map (mtch o apsnd prop_of) (cases ~~ intros)
+ end;
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ ind_cases_setup #>
+ Attrib.setup @{binding mono} (Attrib.add_del mono_add mono_del)
+ "declaration of monotonicity rule";
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = OuterKeyword.keyword "monos";
+
+fun gen_ind_decl mk_def coind =
+ P.fixes -- P.for_fixes --
+ Scan.optional SpecParse.where_alt_specs [] --
+ Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
+ >> (fn (((preds, params), specs), monos) =>
+ (snd oo gen_add_inductive mk_def true coind preds params specs monos));
+
+val ind_decl = gen_ind_decl add_ind_def;
+
+val _ = OuterSyntax.local_theory' "inductive" "define inductive predicates" K.thy_decl (ind_decl false);
+val _ = OuterSyntax.local_theory' "coinductive" "define coinductive predicates" K.thy_decl (ind_decl true);
+
+val _ =
+ OuterSyntax.local_theory "inductive_cases"
+ "create simplified instances of elimination rules (improper)" K.thy_script
+ (P.and_list1 SpecParse.specs >> (snd oo inductive_cases));
+
+end;
+
+end;
--- a/src/HOL/Tools/inductive_codegen.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML Tue Jun 23 15:48:55 2009 +0100
@@ -62,9 +62,9 @@
val nparms = (case optnparms of
SOME k => k
| NONE => (case rules of
- [] => (case try (InductivePackage.the_inductive (ProofContext.init thy)) s of
+ [] => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
SOME (_, {raw_induct, ...}) =>
- length (InductivePackage.params_of raw_induct)
+ length (Inductive.params_of raw_induct)
| NONE => 0)
| xs => snd (snd (snd (split_last xs)))))
in CodegenData.put
@@ -81,11 +81,11 @@
fun get_clauses thy s =
let val {intros, graph, ...} = CodegenData.get thy
in case Symtab.lookup intros s of
- NONE => (case try (InductivePackage.the_inductive (ProofContext.init thy)) s of
+ NONE => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
NONE => NONE
| SOME ({names, ...}, {intrs, raw_induct, ...}) =>
SOME (names, Codegen.thyname_of_const thy s,
- length (InductivePackage.params_of raw_induct),
+ length (Inductive.params_of raw_induct),
preprocess thy intrs))
| SOME _ =>
let
@@ -103,7 +103,7 @@
let
val cnstrs = List.concat (List.concat (map
(map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
- (Symtab.dest (DatatypePackage.get_datatypes thy))));
+ (Symtab.dest (Datatype.get_datatypes thy))));
fun check t = (case strip_comb t of
(Var _, []) => true
| (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of
--- a/src/HOL/Tools/inductive_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,968 +0,0 @@
-(* Title: HOL/Tools/inductive_package.ML
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
- Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
-
-(Co)Inductive Definition module for HOL.
-
-Features:
- * least or greatest fixedpoints
- * mutually recursive definitions
- * definitions involving arbitrary monotone operators
- * automatically proves introduction and elimination rules
-
- Introduction rules have the form
- [| M Pj ti, ..., Q x, ... |] ==> Pk t
- where M is some monotone operator (usually the identity)
- Q x is any side condition on the free variables
- ti, t are any terms
- Pj, Pk are two of the predicates being defined in mutual recursion
-*)
-
-signature BASIC_INDUCTIVE_PACKAGE =
-sig
- type inductive_result
- val morph_result: morphism -> inductive_result -> inductive_result
- type inductive_info
- val the_inductive: Proof.context -> string -> inductive_info
- val print_inductives: Proof.context -> unit
- val mono_add: attribute
- val mono_del: attribute
- val get_monos: Proof.context -> thm list
- val mk_cases: Proof.context -> term -> thm
- val inductive_forall_name: string
- val inductive_forall_def: thm
- val rulify: thm -> thm
- val inductive_cases: (Attrib.binding * string list) list -> local_theory ->
- thm list list * local_theory
- val inductive_cases_i: (Attrib.binding * term list) list -> local_theory ->
- thm list list * local_theory
- type inductive_flags
- val add_inductive_i:
- inductive_flags -> ((binding * typ) * mixfix) list ->
- (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
- inductive_result * local_theory
- val add_inductive: bool -> bool ->
- (binding * string option * mixfix) list ->
- (binding * string option * mixfix) list ->
- (Attrib.binding * string) list ->
- (Facts.ref * Attrib.src list) list ->
- bool -> local_theory -> inductive_result * local_theory
- val add_inductive_global: string -> inductive_flags ->
- ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
- thm list -> theory -> inductive_result * theory
- val arities_of: thm -> (string * int) list
- val params_of: thm -> term list
- val partition_rules: thm -> thm list -> (string * thm list) list
- val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list
- val unpartition_rules: thm list -> (string * 'a list) list -> 'a list
- val infer_intro_vars: thm -> int -> thm list -> term list list
- val setup: theory -> theory
-end;
-
-signature INDUCTIVE_PACKAGE =
-sig
- include BASIC_INDUCTIVE_PACKAGE
- type add_ind_def
- val declare_rules: string -> binding -> bool -> bool -> string list ->
- thm list -> binding list -> Attrib.src list list -> (thm * string list) list ->
- thm -> local_theory -> thm list * thm list * thm * local_theory
- val add_ind_def: add_ind_def
- val gen_add_inductive_i: add_ind_def -> inductive_flags ->
- ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
- thm list -> local_theory -> inductive_result * local_theory
- val gen_add_inductive: add_ind_def -> bool -> bool ->
- (binding * string option * mixfix) list ->
- (binding * string option * mixfix) list ->
- (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
- bool -> local_theory -> inductive_result * local_theory
- val gen_ind_decl: add_ind_def -> bool ->
- OuterParse.token list -> (bool -> local_theory -> local_theory) * OuterParse.token list
-end;
-
-structure InductivePackage: INDUCTIVE_PACKAGE =
-struct
-
-
-(** theory context references **)
-
-val inductive_forall_name = "HOL.induct_forall";
-val inductive_forall_def = thm "induct_forall_def";
-val inductive_conj_name = "HOL.induct_conj";
-val inductive_conj_def = thm "induct_conj_def";
-val inductive_conj = thms "induct_conj";
-val inductive_atomize = thms "induct_atomize";
-val inductive_rulify = thms "induct_rulify";
-val inductive_rulify_fallback = thms "induct_rulify_fallback";
-
-val notTrueE = TrueI RSN (2, notE);
-val notFalseI = Seq.hd (atac 1 notI);
-val simp_thms' = map (fn s => mk_meta_eq (the (find_first
- (equal (OldGoals.read_prop @{theory HOL} s) o prop_of) simp_thms)))
- ["(~True) = False", "(~False) = True",
- "(True --> ?P) = ?P", "(False --> ?P) = True",
- "(?P & True) = ?P", "(True & ?P) = ?P"];
-
-
-
-(** context data **)
-
-type inductive_result =
- {preds: term list, elims: thm list, raw_induct: thm,
- induct: thm, intrs: thm list};
-
-fun morph_result phi {preds, elims, raw_induct: thm, induct, intrs} =
- let
- val term = Morphism.term phi;
- val thm = Morphism.thm phi;
- val fact = Morphism.fact phi;
- in
- {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct,
- induct = thm induct, intrs = fact intrs}
- end;
-
-type inductive_info =
- {names: string list, coind: bool} * inductive_result;
-
-structure InductiveData = GenericDataFun
-(
- type T = inductive_info Symtab.table * thm list;
- val empty = (Symtab.empty, []);
- val extend = I;
- fun merge _ ((tab1, monos1), (tab2, monos2)) =
- (Symtab.merge (K true) (tab1, tab2), Thm.merge_thms (monos1, monos2));
-);
-
-val get_inductives = InductiveData.get o Context.Proof;
-
-fun print_inductives ctxt =
- let
- val (tab, monos) = get_inductives ctxt;
- val space = Consts.space_of (ProofContext.consts_of ctxt);
- in
- [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))),
- Pretty.big_list "monotonicity rules:" (map (ProofContext.pretty_thm ctxt) monos)]
- |> Pretty.chunks |> Pretty.writeln
- end;
-
-
-(* get and put data *)
-
-fun the_inductive ctxt name =
- (case Symtab.lookup (#1 (get_inductives ctxt)) name of
- NONE => error ("Unknown (co)inductive predicate " ^ quote name)
- | SOME info => info);
-
-fun put_inductives names info = InductiveData.map
- (apfst (fold (fn name => Symtab.update (name, info)) names));
-
-
-
-(** monotonicity rules **)
-
-val get_monos = #2 o get_inductives;
-val map_monos = InductiveData.map o apsnd;
-
-fun mk_mono thm =
- let
- val concl = concl_of thm;
- fun eq2mono thm' = [thm' RS (thm' RS eq_to_mono)] @
- (case concl of
- (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
- | _ => [thm' RS (thm' RS eq_to_mono2)]);
- fun dest_less_concl thm = dest_less_concl (thm RS le_funD)
- handle THM _ => thm RS le_boolD
- in
- case concl of
- Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
- | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
- | _ $ (Const ("HOL.ord_class.less_eq", _) $ _ $ _) =>
- [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
- (resolve_tac [le_funI, le_boolI'])) thm))]
- | _ => [thm]
- end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm thm);
-
-val mono_add = Thm.declaration_attribute (map_monos o fold Thm.add_thm o mk_mono);
-val mono_del = Thm.declaration_attribute (map_monos o fold Thm.del_thm o mk_mono);
-
-
-
-(** misc utilities **)
-
-fun message quiet_mode s = if quiet_mode then () else writeln s;
-fun clean_message quiet_mode s = if ! quick_and_dirty then () else message quiet_mode s;
-
-fun coind_prefix true = "co"
- | coind_prefix false = "";
-
-fun log (b:int) m n = if m >= n then 0 else 1 + log b (b * m) n;
-
-fun make_bool_args f g [] i = []
- | make_bool_args f g (x :: xs) i =
- (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2);
-
-fun make_bool_args' xs =
- make_bool_args (K HOLogic.false_const) (K HOLogic.true_const) xs;
-
-fun find_arg T x [] = sys_error "find_arg"
- | find_arg T x ((p as (_, (SOME _, _))) :: ps) =
- apsnd (cons p) (find_arg T x ps)
- | find_arg T x ((p as (U, (NONE, y))) :: ps) =
- if (T: typ) = U then (y, (U, (SOME x, y)) :: ps)
- else apsnd (cons p) (find_arg T x ps);
-
-fun make_args Ts xs =
- map (fn (T, (NONE, ())) => Const (@{const_name undefined}, T) | (_, (SOME t, ())) => t)
- (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts));
-
-fun make_args' Ts xs Us =
- fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs));
-
-fun dest_predicate cs params t =
- let
- val k = length params;
- val (c, ts) = strip_comb t;
- val (xs, ys) = chop k ts;
- val i = find_index_eq c cs;
- in
- if xs = params andalso i >= 0 then
- SOME (c, i, ys, chop (length ys)
- (List.drop (binder_types (fastype_of c), k)))
- else NONE
- end;
-
-fun mk_names a 0 = []
- | mk_names a 1 = [a]
- | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
-
-
-
-(** process rules **)
-
-local
-
-fun err_in_rule ctxt name t msg =
- error (cat_lines ["Ill-formed introduction rule " ^ quote name,
- Syntax.string_of_term ctxt t, msg]);
-
-fun err_in_prem ctxt name t p msg =
- error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p,
- "in introduction rule " ^ quote name, Syntax.string_of_term ctxt t, msg]);
-
-val bad_concl = "Conclusion of introduction rule must be an inductive predicate";
-
-val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate";
-
-val bad_app = "Inductive predicate must be applied to parameter(s) ";
-
-fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
-
-in
-
-fun check_rule ctxt cs params ((binding, att), rule) =
- let
- val err_name = Binding.str_of binding;
- val params' = Term.variant_frees rule (Logic.strip_params rule);
- val frees = rev (map Free params');
- val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
- val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule);
- val rule' = Logic.list_implies (prems, concl);
- val aprems = map (atomize_term (ProofContext.theory_of ctxt)) prems;
- val arule = list_all_free (params', Logic.list_implies (aprems, concl));
-
- fun check_ind err t = case dest_predicate cs params t of
- NONE => err (bad_app ^
- commas (map (Syntax.string_of_term ctxt) params))
- | SOME (_, _, ys, _) =>
- if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs
- then err bad_ind_occ else ();
-
- fun check_prem' prem t =
- if head_of t mem cs then
- check_ind (err_in_prem ctxt err_name rule prem) t
- else (case t of
- Abs (_, _, t) => check_prem' prem t
- | t $ u => (check_prem' prem t; check_prem' prem u)
- | _ => ());
-
- fun check_prem (prem, aprem) =
- if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
- else err_in_prem ctxt err_name rule prem "Non-atomic premise";
- in
- (case concl of
- Const ("Trueprop", _) $ t =>
- if head_of t mem cs then
- (check_ind (err_in_rule ctxt err_name rule') t;
- List.app check_prem (prems ~~ aprems))
- else err_in_rule ctxt err_name rule' bad_concl
- | _ => err_in_rule ctxt err_name rule' bad_concl);
- ((binding, att), arule)
- end;
-
-val rulify =
- hol_simplify inductive_conj
- #> hol_simplify inductive_rulify
- #> hol_simplify inductive_rulify_fallback
- #> Simplifier.norm_hhf;
-
-end;
-
-
-
-(** proofs for (co)inductive predicates **)
-
-(* prove monotonicity *)
-
-fun prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt =
- (message (quiet_mode orelse skip_mono andalso !quick_and_dirty orelse fork_mono)
- " Proving monotonicity ...";
- (if skip_mono then SkipProof.prove else if fork_mono then Goal.prove_future else Goal.prove) ctxt
- [] []
- (HOLogic.mk_Trueprop
- (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
- (fn _ => EVERY [rtac @{thm monoI} 1,
- REPEAT (resolve_tac [le_funI, le_boolI'] 1),
- REPEAT (FIRST
- [atac 1,
- resolve_tac (List.concat (map mk_mono monos) @ get_monos ctxt) 1,
- etac le_funE 1, dtac le_boolD 1])]));
-
-
-(* prove introduction rules *)
-
-fun prove_intrs quiet_mode coind mono fp_def k params intr_ts rec_preds_defs ctxt =
- let
- val _ = clean_message quiet_mode " Proving the introduction rules ...";
-
- val unfold = funpow k (fn th => th RS fun_cong)
- (mono RS (fp_def RS
- (if coind then def_gfp_unfold else def_lfp_unfold)));
-
- fun select_disj 1 1 = []
- | select_disj _ 1 = [rtac disjI1]
- | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
-
- val rules = [refl, TrueI, notFalseI, exI, conjI];
-
- val intrs = map_index (fn (i, intr) => rulify
- (SkipProof.prove ctxt (map (fst o dest_Free) params) [] intr (fn _ => EVERY
- [rewrite_goals_tac rec_preds_defs,
- rtac (unfold RS iffD2) 1,
- EVERY1 (select_disj (length intr_ts) (i + 1)),
- (*Not ares_tac, since refl must be tried before any equality assumptions;
- backtracking may occur if the premises have extra variables!*)
- DEPTH_SOLVE_1 (resolve_tac rules 1 APPEND assume_tac 1)]))) intr_ts
-
- in (intrs, unfold) end;
-
-
-(* prove elimination rules *)
-
-fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt =
- let
- val _ = clean_message quiet_mode " Proving the elimination rules ...";
-
- val ([pname], ctxt') = ctxt |>
- Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
- Variable.variant_fixes ["P"];
- val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
-
- fun dest_intr r =
- (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
- Logic.strip_assums_hyp r, Logic.strip_params r);
-
- val intrs = map dest_intr intr_ts ~~ intr_names;
-
- val rules1 = [disjE, exE, FalseE];
- val rules2 = [conjE, FalseE, notTrueE];
-
- fun prove_elim c =
- let
- val Ts = List.drop (binder_types (fastype_of c), length params);
- val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt';
- val frees = map Free (anames ~~ Ts);
-
- fun mk_elim_prem ((_, _, us, _), ts, params') =
- list_all (params',
- Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
- (frees ~~ us) @ ts, P));
- val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs);
- val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
- map mk_elim_prem (map #1 c_intrs)
- in
- (SkipProof.prove ctxt'' [] prems P
- (fn {prems, ...} => EVERY
- [cut_facts_tac [hd prems] 1,
- rewrite_goals_tac rec_preds_defs,
- dtac (unfold RS iffD1) 1,
- REPEAT (FIRSTGOAL (eresolve_tac rules1)),
- REPEAT (FIRSTGOAL (eresolve_tac rules2)),
- EVERY (map (fn prem =>
- DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_preds_defs prem, conjI] 1)) (tl prems))])
- |> rulify
- |> singleton (ProofContext.export ctxt'' ctxt),
- map #2 c_intrs)
- end
-
- in map prove_elim cs end;
-
-
-(* derivation of simplified elimination rules *)
-
-local
-
-(*delete needless equality assumptions*)
-val refl_thin = Goal.prove_global @{theory HOL} [] [] @{prop "!!P. a = a ==> P ==> P"}
- (fn _ => assume_tac 1);
-val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
-val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
-
-fun simp_case_tac ss i =
- EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i;
-
-in
-
-fun mk_cases ctxt prop =
- let
- val thy = ProofContext.theory_of ctxt;
- val ss = Simplifier.local_simpset_of ctxt;
-
- fun err msg =
- error (Pretty.string_of (Pretty.block
- [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop]));
-
- val elims = Induct.find_casesP ctxt prop;
-
- val cprop = Thm.cterm_of thy prop;
- val tac = ALLGOALS (simp_case_tac ss) THEN prune_params_tac;
- fun mk_elim rl =
- Thm.implies_intr cprop (Tactic.rule_by_tactic tac (Thm.assume cprop RS rl))
- |> singleton (Variable.export (Variable.auto_fixes prop ctxt) ctxt);
- in
- (case get_first (try mk_elim) elims of
- SOME r => r
- | NONE => err "Proposition not an inductive predicate:")
- end;
-
-end;
-
-
-(* inductive_cases *)
-
-fun gen_inductive_cases prep_att prep_prop args lthy =
- let
- val thy = ProofContext.theory_of lthy;
- val facts = args |> map (fn ((a, atts), props) =>
- ((a, map (prep_att thy) atts),
- map (Thm.no_attributes o single o mk_cases lthy o prep_prop lthy) props));
- in lthy |> LocalTheory.notes Thm.generatedK facts |>> map snd end;
-
-val inductive_cases = gen_inductive_cases Attrib.intern_src Syntax.read_prop;
-val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
-
-
-val ind_cases_setup =
- Method.setup @{binding ind_cases}
- (Scan.lift (Scan.repeat1 Args.name_source --
- Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.name) []) >>
- (fn (raw_props, fixes) => fn ctxt =>
- let
- val (_, ctxt') = Variable.add_fixes fixes ctxt;
- val props = Syntax.read_props ctxt' raw_props;
- val ctxt'' = fold Variable.declare_term props ctxt';
- val rules = ProofContext.export ctxt'' ctxt (map (mk_cases ctxt'') props)
- in Method.erule 0 rules end))
- "dynamic case analysis on predicates";
-
-
-(* prove induction rule *)
-
-fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono
- fp_def rec_preds_defs ctxt =
- let
- val _ = clean_message quiet_mode " Proving the induction rule ...";
- val thy = ProofContext.theory_of ctxt;
-
- (* predicates for induction rule *)
-
- val (pnames, ctxt') = ctxt |>
- Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
- Variable.variant_fixes (mk_names "P" (length cs));
- val preds = map Free (pnames ~~
- map (fn c => List.drop (binder_types (fastype_of c), length params) --->
- HOLogic.boolT) cs);
-
- (* transform an introduction rule into a premise for induction rule *)
-
- fun mk_ind_prem r =
- let
- fun subst s = (case dest_predicate cs params s of
- SOME (_, i, ys, (_, Ts)) =>
- let
- val k = length Ts;
- val bs = map Bound (k - 1 downto 0);
- val P = list_comb (List.nth (preds, i),
- map (incr_boundvars k) ys @ bs);
- val Q = list_abs (mk_names "x" k ~~ Ts,
- HOLogic.mk_binop inductive_conj_name
- (list_comb (incr_boundvars k s, bs), P))
- in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end
- | NONE => (case s of
- (t $ u) => (fst (subst t) $ fst (subst u), NONE)
- | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
- | _ => (s, NONE)));
-
- fun mk_prem (s, prems) = (case subst s of
- (_, SOME (t, u)) => t :: u :: prems
- | (t, _) => t :: prems);
-
- val SOME (_, i, ys, _) = dest_predicate cs params
- (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))
-
- in list_all_free (Logic.strip_params r,
- Logic.list_implies (map HOLogic.mk_Trueprop (List.foldr mk_prem
- [] (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r))),
- HOLogic.mk_Trueprop (list_comb (List.nth (preds, i), ys))))
- end;
-
- val ind_prems = map mk_ind_prem intr_ts;
-
-
- (* make conclusions for induction rules *)
-
- val Tss = map (binder_types o fastype_of) preds;
- val (xnames, ctxt'') =
- Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt';
- val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (((xnames, Ts), c), P) =>
- let val frees = map Free (xnames ~~ Ts)
- in HOLogic.mk_imp
- (list_comb (c, params @ frees), list_comb (P, frees))
- end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds)));
-
-
- (* make predicate for instantiation of abstract induction rule *)
-
- val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj
- (map_index (fn (i, P) => List.foldr HOLogic.mk_imp
- (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))
- (make_bool_args HOLogic.mk_not I bs i)) preds));
-
- val ind_concl = HOLogic.mk_Trueprop
- (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred));
-
- val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
-
- val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
- (fn {prems, ...} => EVERY
- [rewrite_goals_tac [inductive_conj_def],
- DETERM (rtac raw_fp_induct 1),
- REPEAT (resolve_tac [le_funI, le_boolI] 1),
- rewrite_goals_tac (inf_fun_eq :: inf_bool_eq :: simp_thms'),
- (*This disjE separates out the introduction rules*)
- REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
- (*Now break down the individual cases. No disjE here in case
- some premise involves disjunction.*)
- REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
- REPEAT (FIRSTGOAL
- (resolve_tac [conjI, impI] ORELSE' (etac notE THEN' atac))),
- EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule
- (inductive_conj_def :: rec_preds_defs @ simp_thms') prem,
- conjI, refl] 1)) prems)]);
-
- val lemma = SkipProof.prove ctxt'' [] []
- (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
- [rewrite_goals_tac rec_preds_defs,
- REPEAT (EVERY
- [REPEAT (resolve_tac [conjI, impI] 1),
- REPEAT (eresolve_tac [le_funE, le_boolE] 1),
- atac 1,
- rewrite_goals_tac simp_thms',
- atac 1])])
-
- in singleton (ProofContext.export ctxt'' ctxt) (induct RS lemma) end;
-
-
-
-(** specification of (co)inductive predicates **)
-
-fun mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts monos params cnames_syn ctxt =
- let
- val fp_name = if coind then @{const_name Inductive.gfp} else @{const_name Inductive.lfp};
-
- val argTs = fold (fn c => fn Ts => Ts @
- (List.drop (binder_types (fastype_of c), length params) \\ Ts)) cs [];
- val k = log 2 1 (length cs);
- val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT;
- val p :: xs = map Free (Variable.variant_frees ctxt intr_ts
- (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs)));
- val bs = map Free (Variable.variant_frees ctxt (p :: xs @ intr_ts)
- (map (rpair HOLogic.boolT) (mk_names "b" k)));
-
- fun subst t = (case dest_predicate cs params t of
- SOME (_, i, ts, (Ts, Us)) =>
- let
- val l = length Us;
- val zs = map Bound (l - 1 downto 0)
- in
- list_abs (map (pair "z") Us, list_comb (p,
- make_bool_args' bs i @ make_args argTs
- ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us))))
- end
- | NONE => (case t of
- t1 $ t2 => subst t1 $ subst t2
- | Abs (x, T, u) => Abs (x, T, subst u)
- | _ => t));
-
- (* transform an introduction rule into a conjunction *)
- (* [| p_i t; ... |] ==> p_j u *)
- (* is transformed into *)
- (* b_j & x_j = u & p b_j t & ... *)
-
- fun transform_rule r =
- let
- val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params
- (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
- val ps = make_bool_args HOLogic.mk_not I bs i @
- map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
- map (subst o HOLogic.dest_Trueprop)
- (Logic.strip_assums_hyp r)
- in List.foldr (fn ((x, T), P) => HOLogic.exists_const T $ (Abs (x, T, P)))
- (if null ps then HOLogic.true_const else foldr1 HOLogic.mk_conj ps)
- (Logic.strip_params r)
- end
-
- (* make a disjunction of all introduction rules *)
-
- val fp_fun = fold_rev lambda (p :: bs @ xs)
- (if null intr_ts then HOLogic.false_const
- else foldr1 HOLogic.mk_disj (map transform_rule intr_ts));
-
- (* add definiton of recursive predicates to theory *)
-
- val rec_name =
- if Binding.is_empty alt_name then
- Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
- else alt_name;
-
- val ((rec_const, (_, fp_def)), ctxt') = ctxt |>
- LocalTheory.define Thm.internalK
- ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn),
- (Attrib.empty_binding, fold_rev lambda params
- (Const (fp_name, (predT --> predT) --> predT) $ fp_fun)));
- val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def])
- (cterm_of (ProofContext.theory_of ctxt') (list_comb (rec_const, params)));
- val specs = if length cs < 2 then [] else
- map_index (fn (i, (name_mx, c)) =>
- let
- val Ts = List.drop (binder_types (fastype_of c), length params);
- val xs = map Free (Variable.variant_frees ctxt intr_ts
- (mk_names "x" (length Ts) ~~ Ts))
- in
- (name_mx, (Attrib.empty_binding, fold_rev lambda (params @ xs)
- (list_comb (rec_const, params @ make_bool_args' bs i @
- make_args argTs (xs ~~ Ts)))))
- end) (cnames_syn ~~ cs);
- val (consts_defs, ctxt'') = fold_map (LocalTheory.define Thm.internalK) specs ctxt';
- val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
-
- val mono = prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt'';
- val ((_, [mono']), ctxt''') =
- LocalTheory.note Thm.internalK (Attrib.empty_binding, [mono]) ctxt'';
-
- in (ctxt''', rec_name, mono', fp_def', map (#2 o #2) consts_defs,
- list_comb (rec_const, params), preds, argTs, bs, xs)
- end;
-
-fun declare_rules kind rec_binding coind no_ind cnames intrs intr_bindings intr_atts
- elims raw_induct ctxt =
- let
- val rec_name = Binding.name_of rec_binding;
- val rec_qualified = Binding.qualify false rec_name;
- val intr_names = map Binding.name_of intr_bindings;
- val ind_case_names = RuleCases.case_names intr_names;
- val induct =
- if coind then
- (raw_induct, [RuleCases.case_names [rec_name],
- RuleCases.case_conclusion (rec_name, intr_names),
- RuleCases.consumes 1, Induct.coinduct_pred (hd cnames)])
- else if no_ind orelse length cnames > 1 then
- (raw_induct, [ind_case_names, RuleCases.consumes 0])
- else (raw_induct RSN (2, rev_mp), [ind_case_names, RuleCases.consumes 1]);
-
- val (intrs', ctxt1) =
- ctxt |>
- LocalTheory.notes kind
- (map rec_qualified intr_bindings ~~ intr_atts ~~ map (fn th => [([th],
- [Attrib.internal (K (ContextRules.intro_query NONE)),
- Attrib.internal (K Nitpick_Ind_Intro_Thms.add)])]) intrs) |>>
- map (hd o snd);
- val (((_, elims'), (_, [induct'])), ctxt2) =
- ctxt1 |>
- LocalTheory.note kind ((rec_qualified (Binding.name "intros"), []), intrs') ||>>
- fold_map (fn (name, (elim, cases)) =>
- LocalTheory.note kind ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "cases"),
- [Attrib.internal (K (RuleCases.case_names cases)),
- Attrib.internal (K (RuleCases.consumes 1)),
- Attrib.internal (K (Induct.cases_pred name)),
- Attrib.internal (K (ContextRules.elim_query NONE))]), [elim]) #>
- apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
- LocalTheory.note kind
- ((rec_qualified (Binding.name (coind_prefix coind ^ "induct")),
- map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]);
-
- val ctxt3 = if no_ind orelse coind then ctxt2 else
- let val inducts = cnames ~~ ProjectRule.projects ctxt2 (1 upto length cnames) induct'
- in
- ctxt2 |>
- LocalTheory.notes kind [((rec_qualified (Binding.name "inducts"), []),
- inducts |> map (fn (name, th) => ([th],
- [Attrib.internal (K ind_case_names),
- Attrib.internal (K (RuleCases.consumes 1)),
- Attrib.internal (K (Induct.induct_pred name))])))] |> snd
- end
- in (intrs', elims', induct', ctxt3) end;
-
-type inductive_flags =
- {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding,
- coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}
-
-type add_ind_def =
- inductive_flags ->
- term list -> (Attrib.binding * term) list -> thm list ->
- term list -> (binding * mixfix) list ->
- local_theory -> inductive_result * local_theory
-
-fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
- cs intros monos params cnames_syn ctxt =
- let
- val _ = null cnames_syn andalso error "No inductive predicates given";
- val names = map (Binding.name_of o fst) cnames_syn;
- val _ = message (quiet_mode andalso not verbose)
- ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names);
-
- val cnames = map (LocalTheory.full_name ctxt o #1) cnames_syn; (* FIXME *)
- val ((intr_names, intr_atts), intr_ts) =
- apfst split_list (split_list (map (check_rule ctxt cs params) intros));
-
- val (ctxt1, rec_name, mono, fp_def, rec_preds_defs, rec_const, preds,
- argTs, bs, xs) = mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts
- monos params cnames_syn ctxt;
-
- val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
- params intr_ts rec_preds_defs ctxt1;
- val elims = if no_elim then [] else
- prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names)
- unfold rec_preds_defs ctxt1;
- val raw_induct = zero_var_indexes
- (if no_ind then Drule.asm_rl else
- if coind then
- singleton (ProofContext.export
- (snd (Variable.add_fixes (map (fst o dest_Free) params) ctxt1)) ctxt1)
- (rotate_prems ~1 (ObjectLogic.rulify
- (fold_rule rec_preds_defs
- (rewrite_rule [le_fun_def, le_bool_def, sup_fun_eq, sup_bool_eq]
- (mono RS (fp_def RS def_coinduct))))))
- else
- prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
- rec_preds_defs ctxt1);
-
- val (intrs', elims', induct, ctxt2) = declare_rules kind rec_name coind no_ind
- cnames intrs intr_names intr_atts elims raw_induct ctxt1;
-
- val result =
- {preds = preds,
- intrs = intrs',
- elims = elims',
- raw_induct = rulify raw_induct,
- induct = induct};
-
- val ctxt3 = ctxt2
- |> LocalTheory.declaration (fn phi =>
- let val result' = morph_result phi result;
- in put_inductives cnames (*global names!?*) ({names = cnames, coind = coind}, result') end);
- in (result, ctxt3) end;
-
-
-(* external interfaces *)
-
-fun gen_add_inductive_i mk_def
- (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono})
- cnames_syn pnames spec monos lthy =
- let
- val thy = ProofContext.theory_of lthy;
- val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
-
-
- (* abbrevs *)
-
- val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy;
-
- fun get_abbrev ((name, atts), t) =
- if can (Logic.strip_assums_concl #> Logic.dest_equals) t then
- let
- val _ = Binding.is_empty name andalso null atts orelse
- error "Abbreviations may not have names or attributes";
- val ((x, T), rhs) = LocalDefs.abs_def (snd (LocalDefs.cert_def ctxt1 t));
- val var =
- (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of
- NONE => error ("Undeclared head of abbreviation " ^ quote x)
- | SOME ((b, T'), mx) =>
- if T <> T' then error ("Bad type specification for abbreviation " ^ quote x)
- else (b, mx));
- in SOME (var, rhs) end
- else NONE;
-
- val abbrevs = map_filter get_abbrev spec;
- val bs = map (Binding.name_of o fst o fst) abbrevs;
-
-
- (* predicates *)
-
- val pre_intros = filter_out (is_some o get_abbrev) spec;
- val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn;
- val cs = map (Free o apfst Binding.name_of o fst) cnames_syn';
- val ps = map Free pnames;
-
- val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn');
- val _ = map (fn abbr => LocalDefs.fixed_abbrev abbr ctxt2) abbrevs;
- val ctxt3 = ctxt2 |> fold (snd oo LocalDefs.fixed_abbrev) abbrevs;
- val expand = Assumption.export_term ctxt3 lthy #> ProofContext.cert_term lthy;
-
- fun close_rule r = list_all_free (rev (fold_aterms
- (fn t as Free (v as (s, _)) =>
- if Variable.is_fixed ctxt1 s orelse
- member (op =) ps t then I else insert (op =) v
- | _ => I) r []), r);
-
- val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros;
- val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn';
- in
- lthy
- |> mk_def flags cs intros monos ps preds
- ||> fold (snd oo LocalTheory.abbrev Syntax.mode_default) abbrevs
- end;
-
-fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy =
- let
- val ((vars, intrs), _) = lthy
- |> ProofContext.set_mode ProofContext.mode_abbrev
- |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs;
- val (cs, ps) = chop (length cnames_syn) vars;
- val monos = Attrib.eval_thms lthy raw_monos;
- val flags = {quiet_mode = false, verbose = verbose, kind = Thm.generatedK,
- alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
- skip_mono = false, fork_mono = not int};
- in
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
- end;
-
-val add_inductive_i = gen_add_inductive_i add_ind_def;
-val add_inductive = gen_add_inductive add_ind_def;
-
-fun add_inductive_global group flags cnames_syn pnames pre_intros monos thy =
- let
- val name = Sign.full_name thy (fst (fst (hd cnames_syn)));
- val ctxt' = thy
- |> TheoryTarget.init NONE
- |> LocalTheory.set_group group
- |> add_inductive_i flags cnames_syn pnames pre_intros monos |> snd
- |> LocalTheory.exit;
- val info = #2 (the_inductive ctxt' name);
- in (info, ProofContext.theory_of ctxt') end;
-
-
-(* read off arities of inductive predicates from raw induction rule *)
-fun arities_of induct =
- map (fn (_ $ t $ u) =>
- (fst (dest_Const (head_of t)), length (snd (strip_comb u))))
- (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
-(* read off parameters of inductive predicate from raw induction rule *)
-fun params_of induct =
- let
- val (_ $ t $ u :: _) =
- HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct));
- val (_, ts) = strip_comb t;
- val (_, us) = strip_comb u
- in
- List.take (ts, length ts - length us)
- end;
-
-val pname_of_intr =
- concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst;
-
-(* partition introduction rules according to predicate name *)
-fun gen_partition_rules f induct intros =
- fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros
- (map (rpair [] o fst) (arities_of induct));
-
-val partition_rules = gen_partition_rules I;
-fun partition_rules' induct = gen_partition_rules fst induct;
-
-fun unpartition_rules intros xs =
- fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r)
- (fn x :: xs => (x, xs)) #>> the) intros xs |> fst;
-
-(* infer order of variables in intro rules from order of quantifiers in elim rule *)
-fun infer_intro_vars elim arity intros =
- let
- val thy = theory_of_thm elim;
- val _ :: cases = prems_of elim;
- val used = map (fst o fst) (Term.add_vars (prop_of elim) []);
- fun mtch (t, u) =
- let
- val params = Logic.strip_params t;
- val vars = map (Var o apfst (rpair 0))
- (Name.variant_list used (map fst params) ~~ map snd params);
- val ts = map (curry subst_bounds (rev vars))
- (List.drop (Logic.strip_assums_hyp t, arity));
- val us = Logic.strip_imp_prems u;
- val tab = fold (Pattern.first_order_match thy) (ts ~~ us)
- (Vartab.empty, Vartab.empty);
- in
- map (Envir.subst_vars tab) vars
- end
- in
- map (mtch o apsnd prop_of) (cases ~~ intros)
- end;
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- ind_cases_setup #>
- Attrib.setup @{binding mono} (Attrib.add_del mono_add mono_del)
- "declaration of monotonicity rule";
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = OuterKeyword.keyword "monos";
-
-fun gen_ind_decl mk_def coind =
- P.fixes -- P.for_fixes --
- Scan.optional SpecParse.where_alt_specs [] --
- Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
- >> (fn (((preds, params), specs), monos) =>
- (snd oo gen_add_inductive mk_def true coind preds params specs monos));
-
-val ind_decl = gen_ind_decl add_ind_def;
-
-val _ = OuterSyntax.local_theory' "inductive" "define inductive predicates" K.thy_decl (ind_decl false);
-val _ = OuterSyntax.local_theory' "coinductive" "define coinductive predicates" K.thy_decl (ind_decl true);
-
-val _ =
- OuterSyntax.local_theory "inductive_cases"
- "create simplified instances of elimination rules (improper)" K.thy_script
- (P.and_list1 SpecParse.specs >> (snd oo inductive_cases));
-
-end;
-
-end;
--- a/src/HOL/Tools/inductive_realizer.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML Tue Jun 23 15:48:55 2009 +0100
@@ -151,7 +151,7 @@
fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
| is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
| is_meta (Const ("Trueprop", _) $ t) = (case head_of t of
- Const (s, _) => can (InductivePackage.the_inductive ctxt) s
+ Const (s, _) => can (Inductive.the_inductive ctxt) s
| _ => true)
| is_meta _ = false;
@@ -277,13 +277,13 @@
val inducts = PureThy.get_thms thy (Long_Name.qualify qualifier "inducts");
val iTs = OldTerm.term_tvars (prop_of (hd intrs));
val ar = length vs + length iTs;
- val params = InductivePackage.params_of raw_induct;
- val arities = InductivePackage.arities_of raw_induct;
+ val params = Inductive.params_of raw_induct;
+ val arities = Inductive.arities_of raw_induct;
val nparms = length params;
val params' = map dest_Var params;
- val rss = InductivePackage.partition_rules raw_induct intrs;
+ val rss = Inductive.partition_rules raw_induct intrs;
val rss' = map (fn (((s, rs), (_, arity)), elim) =>
- (s, (InductivePackage.infer_intro_vars elim arity rs ~~ rs)))
+ (s, (Inductive.infer_intro_vars elim arity rs ~~ rs)))
(rss ~~ arities ~~ elims);
val (prfx, _) = split_last (Long_Name.explode (fst (hd rss)));
val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets;
@@ -307,7 +307,7 @@
val ((dummies, dt_info), thy2) =
thy1
- |> add_dummies (DatatypePackage.add_datatype
+ |> add_dummies (Datatype.add_datatype
{ strict = false, flat_names = false, quiet = false } (map (Binding.name_of o #2) dts))
(map (pair false) dts) []
||> Extraction.add_typeof_eqns_i ty_eqs
@@ -348,7 +348,7 @@
(** realizability predicate **)
val (ind_info, thy3') = thy2 |>
- InductivePackage.add_inductive_global (serial_string ())
+ Inductive.add_inductive_global (serial_string ())
{quiet_mode = false, verbose = false, kind = Thm.generatedK, alt_name = Binding.empty,
coind = false, no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
rlzpreds rlzparams (map (fn (rintr, intr) =>
@@ -483,7 +483,7 @@
fun add_ind_realizers name rsets thy =
let
val (_, {intrs, induct, raw_induct, elims, ...}) =
- InductivePackage.the_inductive (ProofContext.init thy) name;
+ Inductive.the_inductive (ProofContext.init thy) name;
val vss = sort (int_ord o pairself length)
(subsets (map fst (relevant_vars (concl_of (hd intrs)))))
in
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/inductive_set.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,566 @@
+(* Title: HOL/Tools/inductive_set.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Wrapper for defining inductive sets using package for inductive predicates,
+including infrastructure for converting between predicates and sets.
+*)
+
+signature INDUCTIVE_SET =
+sig
+ val to_set_att: thm list -> attribute
+ val to_pred_att: thm list -> attribute
+ val pred_set_conv_att: attribute
+ val add_inductive_i:
+ Inductive.inductive_flags ->
+ ((binding * typ) * mixfix) list ->
+ (string * typ) list ->
+ (Attrib.binding * term) list -> thm list ->
+ local_theory -> Inductive.inductive_result * local_theory
+ val add_inductive: bool -> bool ->
+ (binding * string option * mixfix) list ->
+ (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
+ bool -> local_theory -> Inductive.inductive_result * local_theory
+ val codegen_preproc: theory -> thm list -> thm list
+ val setup: theory -> theory
+end;
+
+structure Inductive_Set: INDUCTIVE_SET =
+struct
+
+(**** simplify {(x1, ..., xn). (x1, ..., xn) : S} to S ****)
+
+val collect_mem_simproc =
+ Simplifier.simproc (theory "Set") "Collect_mem" ["Collect t"] (fn thy => fn ss =>
+ fn S as Const ("Collect", Type ("fun", [_, T])) $ t =>
+ let val (u, Ts, ps) = HOLogic.strip_split t
+ in case u of
+ (c as Const ("op :", _)) $ q $ S' =>
+ (case try (HOLogic.dest_tuple' ps) q of
+ NONE => NONE
+ | SOME ts =>
+ if not (loose_bvar (S', 0)) andalso
+ ts = map Bound (length ps downto 0)
+ then
+ let val simp = full_simp_tac (Simplifier.inherit_context ss
+ (HOL_basic_ss addsimps [split_paired_all, split_conv])) 1
+ in
+ SOME (Goal.prove (Simplifier.the_context ss) [] []
+ (Const ("==", T --> T --> propT) $ S $ S')
+ (K (EVERY
+ [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
+ rtac subsetI 1, dtac CollectD 1, simp,
+ rtac subsetI 1, rtac CollectI 1, simp])))
+ end
+ else NONE)
+ | _ => NONE
+ end
+ | _ => NONE);
+
+(***********************************************************************************)
+(* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *)
+(* and (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y}) *)
+(* used for converting "strong" (co)induction rules *)
+(***********************************************************************************)
+
+val anyt = Free ("t", TFree ("'t", []));
+
+fun strong_ind_simproc tab =
+ Simplifier.simproc_i @{theory HOL} "strong_ind" [anyt] (fn thy => fn ss => fn t =>
+ let
+ fun close p t f =
+ let val vs = Term.add_vars t []
+ in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
+ (p (fold (Logic.all o Var) vs t) f)
+ end;
+ fun mkop "op &" T x = SOME (Const (@{const_name "Int"}, T --> T --> T), x)
+ | mkop "op |" T x = SOME (Const (@{const_name "Un"}, T --> T --> T), x)
+ | mkop _ _ _ = NONE;
+ fun mk_collect p T t =
+ let val U = HOLogic.dest_setT T
+ in HOLogic.Collect_const U $
+ HOLogic.ap_split' (HOLogic.prod_factors p) U HOLogic.boolT t
+ end;
+ fun decomp (Const (s, _) $ ((m as Const ("op :",
+ Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) =
+ mkop s T (m, p, S, mk_collect p T (head_of u))
+ | decomp (Const (s, _) $ u $ ((m as Const ("op :",
+ Type (_, [_, Type (_, [T, _])]))) $ p $ S)) =
+ mkop s T (m, p, mk_collect p T (head_of u), S)
+ | decomp _ = NONE;
+ val simp = full_simp_tac (Simplifier.inherit_context ss
+ (HOL_basic_ss addsimps [mem_Collect_eq, split_conv])) 1;
+ fun mk_rew t = (case strip_abs_vars t of
+ [] => NONE
+ | xs => (case decomp (strip_abs_body t) of
+ NONE => NONE
+ | SOME (bop, (m, p, S, S')) =>
+ SOME (close (Goal.prove (Simplifier.the_context ss) [] [])
+ (Logic.mk_equals (t, list_abs (xs, m $ p $ (bop $ S $ S'))))
+ (K (EVERY
+ [rtac eq_reflection 1, REPEAT (rtac ext 1), rtac iffI 1,
+ EVERY [etac conjE 1, rtac IntI 1, simp, simp,
+ etac IntE 1, rtac conjI 1, simp, simp] ORELSE
+ EVERY [etac disjE 1, rtac UnI1 1, simp, rtac UnI2 1, simp,
+ etac UnE 1, rtac disjI1 1, simp, rtac disjI2 1, simp]])))
+ handle ERROR _ => NONE))
+ in
+ case strip_comb t of
+ (h as Const (name, _), ts) => (case Symtab.lookup tab name of
+ SOME _ =>
+ let val rews = map mk_rew ts
+ in
+ if forall is_none rews then NONE
+ else SOME (fold (fn th1 => fn th2 => combination th2 th1)
+ (map2 (fn SOME r => K r | NONE => reflexive o cterm_of thy)
+ rews ts) (reflexive (cterm_of thy h)))
+ end
+ | NONE => NONE)
+ | _ => NONE
+ end);
+
+(* only eta contract terms occurring as arguments of functions satisfying p *)
+fun eta_contract p =
+ let
+ fun eta b (Abs (a, T, body)) =
+ (case eta b body of
+ body' as (f $ Bound 0) =>
+ if loose_bvar1 (f, 0) orelse not b then Abs (a, T, body')
+ else incr_boundvars ~1 f
+ | body' => Abs (a, T, body'))
+ | eta b (t $ u) = eta b t $ eta (p (head_of t)) u
+ | eta b t = t
+ in eta false end;
+
+fun eta_contract_thm p =
+ Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct =>
+ Thm.transitive (Thm.eta_conversion ct)
+ (Thm.symmetric (Thm.eta_conversion
+ (cterm_of (theory_of_cterm ct) (eta_contract p (term_of ct)))))));
+
+
+(***********************************************************)
+(* rules for converting between predicate and set notation *)
+(* *)
+(* rules for converting predicates to sets have the form *)
+(* P (%x y. (x, y) : s) = (%x y. (x, y) : S s) *)
+(* *)
+(* rules for converting sets to predicates have the form *)
+(* S {(x, y). p x y} = {(x, y). P p x y} *)
+(* *)
+(* where s and p are parameters *)
+(***********************************************************)
+
+structure PredSetConvData = GenericDataFun
+(
+ type T =
+ {(* rules for converting predicates to sets *)
+ to_set_simps: thm list,
+ (* rules for converting sets to predicates *)
+ to_pred_simps: thm list,
+ (* arities of functions of type t set => ... => u set *)
+ set_arities: (typ * (int list list option list * int list list option)) list Symtab.table,
+ (* arities of functions of type (t => ... => bool) => u => ... => bool *)
+ pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table};
+ val empty = {to_set_simps = [], to_pred_simps = [],
+ set_arities = Symtab.empty, pred_arities = Symtab.empty};
+ val extend = I;
+ fun merge _
+ ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1,
+ set_arities = set_arities1, pred_arities = pred_arities1},
+ {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2,
+ set_arities = set_arities2, pred_arities = pred_arities2}) : T =
+ {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2),
+ to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2),
+ set_arities = Symtab.merge_list op = (set_arities1, set_arities2),
+ pred_arities = Symtab.merge_list op = (pred_arities1, pred_arities2)};
+);
+
+fun name_type_of (Free p) = SOME p
+ | name_type_of (Const p) = SOME p
+ | name_type_of _ = NONE;
+
+fun map_type f (Free (s, T)) = Free (s, f T)
+ | map_type f (Var (ixn, T)) = Var (ixn, f T)
+ | map_type f _ = error "map_type";
+
+fun find_most_specific is_inst f eq xs T =
+ find_first (fn U => is_inst (T, f U)
+ andalso forall (fn U' => eq (f U, f U') orelse not
+ (is_inst (T, f U') andalso is_inst (f U', f U)))
+ xs) xs;
+
+fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of
+ NONE => NONE
+ | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T;
+
+fun lookup_rule thy f rules = find_most_specific
+ (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules;
+
+fun infer_arities thy arities (optf, t) fs = case strip_comb t of
+ (Abs (s, T, u), []) => infer_arities thy arities (NONE, u) fs
+ | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs
+ | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of
+ SOME (SOME (_, (arity, _))) =>
+ (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs
+ handle Subscript => error "infer_arities: bad term")
+ | _ => fold (infer_arities thy arities) (map (pair NONE) ts)
+ (case optf of
+ NONE => fs
+ | SOME f => AList.update op = (u, the_default f
+ (Option.map (curry op inter f) (AList.lookup op = fs u))) fs));
+
+
+(**************************************************************)
+(* derive the to_pred equation from the to_set equation *)
+(* *)
+(* 1. instantiate each set parameter with {(x, y). p x y} *)
+(* 2. apply %P. {(x, y). P x y} to both sides of the equation *)
+(* 3. simplify *)
+(**************************************************************)
+
+fun mk_to_pred_inst thy fs =
+ map (fn (x, ps) =>
+ let
+ val U = HOLogic.dest_setT (fastype_of x);
+ val x' = map_type (K (HOLogic.prodT_factors' ps U ---> HOLogic.boolT)) x
+ in
+ (cterm_of thy x,
+ cterm_of thy (HOLogic.Collect_const U $
+ HOLogic.ap_split' ps U HOLogic.boolT x'))
+ end) fs;
+
+fun mk_to_pred_eq p fs optfs' T thm =
+ let
+ val thy = theory_of_thm thm;
+ val insts = mk_to_pred_inst thy fs;
+ val thm' = Thm.instantiate ([], insts) thm;
+ val thm'' = (case optfs' of
+ NONE => thm' RS sym
+ | SOME fs' =>
+ let
+ val (_, U) = split_last (binder_types T);
+ val Ts = HOLogic.prodT_factors' fs' U;
+ (* FIXME: should cterm_instantiate increment indexes? *)
+ val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
+ val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
+ Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
+ in
+ thm' RS (Drule.cterm_instantiate [(arg_cong_f,
+ cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT,
+ HOLogic.Collect_const U $ HOLogic.ap_split' fs' U
+ HOLogic.boolT (Bound 0))))] arg_cong' RS sym)
+ end)
+ in
+ Simplifier.simplify (HOL_basic_ss addsimps [mem_Collect_eq, split_conv]
+ addsimprocs [collect_mem_simproc]) thm'' |>
+ zero_var_indexes |> eta_contract_thm (equal p)
+ end;
+
+
+(**** declare rules for converting predicates to sets ****)
+
+fun add ctxt thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
+ case prop_of thm of
+ Const ("Trueprop", _) $ (Const ("op =", Type (_, [T, _])) $ lhs $ rhs) =>
+ (case body_type T of
+ Type ("bool", []) =>
+ let
+ val thy = Context.theory_of ctxt;
+ fun factors_of t fs = case strip_abs_body t of
+ Const ("op :", _) $ u $ S =>
+ if is_Free S orelse is_Var S then
+ let val ps = HOLogic.prod_factors u
+ in (SOME ps, (S, ps) :: fs) end
+ else (NONE, fs)
+ | _ => (NONE, fs);
+ val (h, ts) = strip_comb lhs
+ val (pfs, fs) = fold_map factors_of ts [];
+ val ((h', ts'), fs') = (case rhs of
+ Abs _ => (case strip_abs_body rhs of
+ Const ("op :", _) $ u $ S =>
+ (strip_comb S, SOME (HOLogic.prod_factors u))
+ | _ => error "member symbol on right-hand side expected")
+ | _ => (strip_comb rhs, NONE))
+ in
+ case (name_type_of h, name_type_of h') of
+ (SOME (s, T), SOME (s', T')) =>
+ if exists (fn (U, _) =>
+ Sign.typ_instance thy (T', U) andalso
+ Sign.typ_instance thy (U, T'))
+ (Symtab.lookup_list set_arities s')
+ then
+ (warning ("Ignoring conversion rule for operator " ^ s'); tab)
+ else
+ {to_set_simps = thm :: to_set_simps,
+ to_pred_simps =
+ mk_to_pred_eq h fs fs' T' thm :: to_pred_simps,
+ set_arities = Symtab.insert_list op = (s',
+ (T', (map (AList.lookup op = fs) ts', fs'))) set_arities,
+ pred_arities = Symtab.insert_list op = (s,
+ (T, (pfs, fs'))) pred_arities}
+ | _ => error "set / predicate constant expected"
+ end
+ | _ => error "equation between predicates expected")
+ | _ => error "equation expected";
+
+val pred_set_conv_att = Thm.declaration_attribute
+ (fn thm => fn ctxt => PredSetConvData.map (add ctxt thm) ctxt);
+
+
+(**** convert theorem in set notation to predicate notation ****)
+
+fun is_pred tab t =
+ case Option.map (Symtab.lookup tab o fst) (name_type_of t) of
+ SOME (SOME _) => true | _ => false;
+
+fun to_pred_simproc rules =
+ let val rules' = map mk_meta_eq rules
+ in
+ Simplifier.simproc_i @{theory HOL} "to_pred" [anyt]
+ (fn thy => K (lookup_rule thy (prop_of #> Logic.dest_equals) rules'))
+ end;
+
+fun to_pred_proc thy rules t = case lookup_rule thy I rules t of
+ NONE => NONE
+ | SOME (lhs, rhs) =>
+ SOME (Envir.subst_vars
+ (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs);
+
+fun to_pred thms ctxt thm =
+ let
+ val thy = Context.theory_of ctxt;
+ val {to_pred_simps, set_arities, pred_arities, ...} =
+ fold (add ctxt) thms (PredSetConvData.get ctxt);
+ val fs = filter (is_Var o fst)
+ (infer_arities thy set_arities (NONE, prop_of thm) []);
+ (* instantiate each set parameter with {(x, y). p x y} *)
+ val insts = mk_to_pred_inst thy fs
+ in
+ thm |>
+ Thm.instantiate ([], insts) |>
+ Simplifier.full_simplify (HOL_basic_ss addsimprocs
+ [to_pred_simproc (mem_Collect_eq :: split_conv :: to_pred_simps)]) |>
+ eta_contract_thm (is_pred pred_arities) |>
+ RuleCases.save thm
+ end;
+
+val to_pred_att = Thm.rule_attribute o to_pred;
+
+
+(**** convert theorem in predicate notation to set notation ****)
+
+fun to_set thms ctxt thm =
+ let
+ val thy = Context.theory_of ctxt;
+ val {to_set_simps, pred_arities, ...} =
+ fold (add ctxt) thms (PredSetConvData.get ctxt);
+ val fs = filter (is_Var o fst)
+ (infer_arities thy pred_arities (NONE, prop_of thm) []);
+ (* instantiate each predicate parameter with %x y. (x, y) : s *)
+ val insts = map (fn (x, ps) =>
+ let
+ val Ts = binder_types (fastype_of x);
+ val T = HOLogic.mk_tupleT ps Ts;
+ val x' = map_type (K (HOLogic.mk_setT T)) x
+ in
+ (cterm_of thy x,
+ cterm_of thy (list_abs (map (pair "x") Ts, HOLogic.mk_mem
+ (HOLogic.mk_tuple' ps T (map Bound (length ps downto 0)), x'))))
+ end) fs
+ in
+ thm |>
+ Thm.instantiate ([], insts) |>
+ Simplifier.full_simplify (HOL_basic_ss addsimps to_set_simps
+ addsimprocs [strong_ind_simproc pred_arities, collect_mem_simproc]) |>
+ RuleCases.save thm
+ end;
+
+val to_set_att = Thm.rule_attribute o to_set;
+
+
+(**** preprocessor for code generator ****)
+
+fun codegen_preproc thy =
+ let
+ val {to_pred_simps, set_arities, pred_arities, ...} =
+ PredSetConvData.get (Context.Theory thy);
+ fun preproc thm =
+ if exists_Const (fn (s, _) => case Symtab.lookup set_arities s of
+ NONE => false
+ | SOME arities => exists (fn (_, (xs, _)) =>
+ forall is_none xs) arities) (prop_of thm)
+ then
+ thm |>
+ Simplifier.full_simplify (HOL_basic_ss addsimprocs
+ [to_pred_simproc (mem_Collect_eq :: split_conv :: to_pred_simps)]) |>
+ eta_contract_thm (is_pred pred_arities)
+ else thm
+ in map preproc end;
+
+fun code_ind_att optmod = to_pred_att [] #> InductiveCodegen.add optmod NONE;
+
+
+(**** definition of inductive sets ****)
+
+fun add_ind_set_def
+ {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
+ cs intros monos params cnames_syn ctxt =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val {set_arities, pred_arities, to_pred_simps, ...} =
+ PredSetConvData.get (Context.Proof ctxt);
+ fun infer (Abs (_, _, t)) = infer t
+ | infer (Const ("op :", _) $ t $ u) =
+ infer_arities thy set_arities (SOME (HOLogic.prod_factors t), u)
+ | infer (t $ u) = infer t #> infer u
+ | infer _ = I;
+ val new_arities = filter_out
+ (fn (x as Free (_, T), _) => x mem params andalso length (binder_types T) > 1
+ | _ => false) (fold (snd #> infer) intros []);
+ val params' = map (fn x => (case AList.lookup op = new_arities x of
+ SOME fs =>
+ let
+ val T = HOLogic.dest_setT (fastype_of x);
+ val Ts = HOLogic.prodT_factors' fs T;
+ val x' = map_type (K (Ts ---> HOLogic.boolT)) x
+ in
+ (x, (x',
+ (HOLogic.Collect_const T $
+ HOLogic.ap_split' fs T HOLogic.boolT x',
+ list_abs (map (pair "x") Ts, HOLogic.mk_mem
+ (HOLogic.mk_tuple' fs T (map Bound (length fs downto 0)),
+ x)))))
+ end
+ | NONE => (x, (x, (x, x))))) params;
+ val (params1, (params2, params3)) =
+ params' |> map snd |> split_list ||> split_list;
+ val paramTs = map fastype_of params;
+
+ (* equations for converting sets to predicates *)
+ val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) =>
+ let
+ val fs = the_default [] (AList.lookup op = new_arities c);
+ val (Us, U) = split_last (binder_types T);
+ val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks
+ [Pretty.str "Argument types",
+ Pretty.block (Pretty.commas (map (Syntax.pretty_typ ctxt) Us)),
+ Pretty.str ("of " ^ s ^ " do not agree with types"),
+ Pretty.block (Pretty.commas (map (Syntax.pretty_typ ctxt) paramTs)),
+ Pretty.str "of declared parameters"]));
+ val Ts = HOLogic.prodT_factors' fs U;
+ val c' = Free (s ^ "p",
+ map fastype_of params1 @ Ts ---> HOLogic.boolT)
+ in
+ ((c', (fs, U, Ts)),
+ (list_comb (c, params2),
+ HOLogic.Collect_const U $ HOLogic.ap_split' fs U HOLogic.boolT
+ (list_comb (c', params1))))
+ end) |> split_list |>> split_list;
+ val eqns' = eqns @
+ map (prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq)
+ (mem_Collect_eq :: split_conv :: to_pred_simps);
+
+ (* predicate version of the introduction rules *)
+ val intros' =
+ map (fn (name_atts, t) => (name_atts,
+ t |>
+ map_aterms (fn u =>
+ (case AList.lookup op = params' u of
+ SOME (_, (u', _)) => u'
+ | NONE => u)) |>
+ Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |>
+ eta_contract (member op = cs' orf is_pred pred_arities))) intros;
+ val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn;
+ val monos' = map (to_pred [] (Context.Proof ctxt)) monos;
+ val ({preds, intrs, elims, raw_induct, ...}, ctxt1) =
+ Inductive.add_ind_def
+ {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Binding.empty,
+ coind = coind, no_elim = no_elim, no_ind = no_ind,
+ skip_mono = skip_mono, fork_mono = fork_mono}
+ cs' intros' monos' params1 cnames_syn' ctxt;
+
+ (* define inductive sets using previously defined predicates *)
+ val (defs, ctxt2) = fold_map (LocalTheory.define Thm.internalK)
+ (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.empty_binding,
+ fold_rev lambda params (HOLogic.Collect_const U $
+ HOLogic.ap_split' fs U HOLogic.boolT (list_comb (p, params3))))))
+ (cnames_syn ~~ cs_info ~~ preds)) ctxt1;
+
+ (* prove theorems for converting predicate to set notation *)
+ val ctxt3 = fold
+ (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn ctxt =>
+ let val conv_thm =
+ Goal.prove ctxt (map (fst o dest_Free) params) []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (list_comb (p, params3),
+ list_abs (map (pair "x") Ts, HOLogic.mk_mem
+ (HOLogic.mk_tuple' fs U (map Bound (length fs downto 0)),
+ list_comb (c, params))))))
+ (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps
+ [def, mem_Collect_eq, split_conv]) 1))
+ in
+ ctxt |> LocalTheory.note kind ((Binding.name (s ^ "p_" ^ s ^ "_eq"),
+ [Attrib.internal (K pred_set_conv_att)]),
+ [conv_thm]) |> snd
+ end) (preds ~~ cs ~~ cs_info ~~ defs) ctxt2;
+
+ (* convert theorems to set notation *)
+ val rec_name =
+ if Binding.is_empty alt_name then
+ Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
+ else alt_name;
+ val cnames = map (LocalTheory.full_name ctxt3 o #1) cnames_syn; (* FIXME *)
+ val (intr_names, intr_atts) = split_list (map fst intros);
+ val raw_induct' = to_set [] (Context.Proof ctxt3) raw_induct;
+ val (intrs', elims', induct, ctxt4) =
+ Inductive.declare_rules kind rec_name coind no_ind cnames
+ (map (to_set [] (Context.Proof ctxt3)) intrs) intr_names intr_atts
+ (map (fn th => (to_set [] (Context.Proof ctxt3) th,
+ map fst (fst (RuleCases.get th)))) elims)
+ raw_induct' ctxt3
+ in
+ ({intrs = intrs', elims = elims', induct = induct,
+ raw_induct = raw_induct', preds = map fst defs},
+ ctxt4)
+ end;
+
+val add_inductive_i = Inductive.gen_add_inductive_i add_ind_set_def;
+val add_inductive = Inductive.gen_add_inductive add_ind_set_def;
+
+val mono_add_att = to_pred_att [] #> Inductive.mono_add;
+val mono_del_att = to_pred_att [] #> Inductive.mono_del;
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ Attrib.setup @{binding pred_set_conv} (Scan.succeed pred_set_conv_att)
+ "declare rules for converting between predicate and set notation" #>
+ Attrib.setup @{binding to_set} (Attrib.thms >> to_set_att) "convert rule to set notation" #>
+ Attrib.setup @{binding to_pred} (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #>
+ Code.add_attribute ("ind_set",
+ Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> code_ind_att) #>
+ Codegen.add_preprocessor codegen_preproc #>
+ Attrib.setup @{binding mono_set} (Attrib.add_del mono_add_att mono_del_att)
+ "declaration of monotonicity rule for set operators" #>
+ Context.theory_map (Simplifier.map_ss (fn ss => ss addsimprocs [collect_mem_simproc]));
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def;
+
+val _ =
+ OuterSyntax.local_theory' "inductive_set" "define inductive sets" K.thy_decl (ind_set_decl false);
+
+val _ =
+ OuterSyntax.local_theory' "coinductive_set" "define coinductive sets" K.thy_decl (ind_set_decl true);
+
+end;
+
+end;
--- a/src/HOL/Tools/inductive_set_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,566 +0,0 @@
-(* Title: HOL/Tools/inductive_set_package.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Wrapper for defining inductive sets using package for inductive predicates,
-including infrastructure for converting between predicates and sets.
-*)
-
-signature INDUCTIVE_SET_PACKAGE =
-sig
- val to_set_att: thm list -> attribute
- val to_pred_att: thm list -> attribute
- val pred_set_conv_att: attribute
- val add_inductive_i:
- InductivePackage.inductive_flags ->
- ((binding * typ) * mixfix) list ->
- (string * typ) list ->
- (Attrib.binding * term) list -> thm list ->
- local_theory -> InductivePackage.inductive_result * local_theory
- val add_inductive: bool -> bool ->
- (binding * string option * mixfix) list ->
- (binding * string option * mixfix) list ->
- (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
- bool -> local_theory -> InductivePackage.inductive_result * local_theory
- val codegen_preproc: theory -> thm list -> thm list
- val setup: theory -> theory
-end;
-
-structure InductiveSetPackage: INDUCTIVE_SET_PACKAGE =
-struct
-
-(**** simplify {(x1, ..., xn). (x1, ..., xn) : S} to S ****)
-
-val collect_mem_simproc =
- Simplifier.simproc (theory "Set") "Collect_mem" ["Collect t"] (fn thy => fn ss =>
- fn S as Const ("Collect", Type ("fun", [_, T])) $ t =>
- let val (u, Ts, ps) = HOLogic.strip_split t
- in case u of
- (c as Const ("op :", _)) $ q $ S' =>
- (case try (HOLogic.dest_tuple' ps) q of
- NONE => NONE
- | SOME ts =>
- if not (loose_bvar (S', 0)) andalso
- ts = map Bound (length ps downto 0)
- then
- let val simp = full_simp_tac (Simplifier.inherit_context ss
- (HOL_basic_ss addsimps [split_paired_all, split_conv])) 1
- in
- SOME (Goal.prove (Simplifier.the_context ss) [] []
- (Const ("==", T --> T --> propT) $ S $ S')
- (K (EVERY
- [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
- rtac subsetI 1, dtac CollectD 1, simp,
- rtac subsetI 1, rtac CollectI 1, simp])))
- end
- else NONE)
- | _ => NONE
- end
- | _ => NONE);
-
-(***********************************************************************************)
-(* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *)
-(* and (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y}) *)
-(* used for converting "strong" (co)induction rules *)
-(***********************************************************************************)
-
-val anyt = Free ("t", TFree ("'t", []));
-
-fun strong_ind_simproc tab =
- Simplifier.simproc_i @{theory HOL} "strong_ind" [anyt] (fn thy => fn ss => fn t =>
- let
- fun close p t f =
- let val vs = Term.add_vars t []
- in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
- (p (fold (Logic.all o Var) vs t) f)
- end;
- fun mkop "op &" T x = SOME (Const (@{const_name "Int"}, T --> T --> T), x)
- | mkop "op |" T x = SOME (Const (@{const_name "Un"}, T --> T --> T), x)
- | mkop _ _ _ = NONE;
- fun mk_collect p T t =
- let val U = HOLogic.dest_setT T
- in HOLogic.Collect_const U $
- HOLogic.ap_split' (HOLogic.prod_factors p) U HOLogic.boolT t
- end;
- fun decomp (Const (s, _) $ ((m as Const ("op :",
- Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) =
- mkop s T (m, p, S, mk_collect p T (head_of u))
- | decomp (Const (s, _) $ u $ ((m as Const ("op :",
- Type (_, [_, Type (_, [T, _])]))) $ p $ S)) =
- mkop s T (m, p, mk_collect p T (head_of u), S)
- | decomp _ = NONE;
- val simp = full_simp_tac (Simplifier.inherit_context ss
- (HOL_basic_ss addsimps [mem_Collect_eq, split_conv])) 1;
- fun mk_rew t = (case strip_abs_vars t of
- [] => NONE
- | xs => (case decomp (strip_abs_body t) of
- NONE => NONE
- | SOME (bop, (m, p, S, S')) =>
- SOME (close (Goal.prove (Simplifier.the_context ss) [] [])
- (Logic.mk_equals (t, list_abs (xs, m $ p $ (bop $ S $ S'))))
- (K (EVERY
- [rtac eq_reflection 1, REPEAT (rtac ext 1), rtac iffI 1,
- EVERY [etac conjE 1, rtac IntI 1, simp, simp,
- etac IntE 1, rtac conjI 1, simp, simp] ORELSE
- EVERY [etac disjE 1, rtac UnI1 1, simp, rtac UnI2 1, simp,
- etac UnE 1, rtac disjI1 1, simp, rtac disjI2 1, simp]])))
- handle ERROR _ => NONE))
- in
- case strip_comb t of
- (h as Const (name, _), ts) => (case Symtab.lookup tab name of
- SOME _ =>
- let val rews = map mk_rew ts
- in
- if forall is_none rews then NONE
- else SOME (fold (fn th1 => fn th2 => combination th2 th1)
- (map2 (fn SOME r => K r | NONE => reflexive o cterm_of thy)
- rews ts) (reflexive (cterm_of thy h)))
- end
- | NONE => NONE)
- | _ => NONE
- end);
-
-(* only eta contract terms occurring as arguments of functions satisfying p *)
-fun eta_contract p =
- let
- fun eta b (Abs (a, T, body)) =
- (case eta b body of
- body' as (f $ Bound 0) =>
- if loose_bvar1 (f, 0) orelse not b then Abs (a, T, body')
- else incr_boundvars ~1 f
- | body' => Abs (a, T, body'))
- | eta b (t $ u) = eta b t $ eta (p (head_of t)) u
- | eta b t = t
- in eta false end;
-
-fun eta_contract_thm p =
- Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct =>
- Thm.transitive (Thm.eta_conversion ct)
- (Thm.symmetric (Thm.eta_conversion
- (cterm_of (theory_of_cterm ct) (eta_contract p (term_of ct)))))));
-
-
-(***********************************************************)
-(* rules for converting between predicate and set notation *)
-(* *)
-(* rules for converting predicates to sets have the form *)
-(* P (%x y. (x, y) : s) = (%x y. (x, y) : S s) *)
-(* *)
-(* rules for converting sets to predicates have the form *)
-(* S {(x, y). p x y} = {(x, y). P p x y} *)
-(* *)
-(* where s and p are parameters *)
-(***********************************************************)
-
-structure PredSetConvData = GenericDataFun
-(
- type T =
- {(* rules for converting predicates to sets *)
- to_set_simps: thm list,
- (* rules for converting sets to predicates *)
- to_pred_simps: thm list,
- (* arities of functions of type t set => ... => u set *)
- set_arities: (typ * (int list list option list * int list list option)) list Symtab.table,
- (* arities of functions of type (t => ... => bool) => u => ... => bool *)
- pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table};
- val empty = {to_set_simps = [], to_pred_simps = [],
- set_arities = Symtab.empty, pred_arities = Symtab.empty};
- val extend = I;
- fun merge _
- ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1,
- set_arities = set_arities1, pred_arities = pred_arities1},
- {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2,
- set_arities = set_arities2, pred_arities = pred_arities2}) : T =
- {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2),
- to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2),
- set_arities = Symtab.merge_list op = (set_arities1, set_arities2),
- pred_arities = Symtab.merge_list op = (pred_arities1, pred_arities2)};
-);
-
-fun name_type_of (Free p) = SOME p
- | name_type_of (Const p) = SOME p
- | name_type_of _ = NONE;
-
-fun map_type f (Free (s, T)) = Free (s, f T)
- | map_type f (Var (ixn, T)) = Var (ixn, f T)
- | map_type f _ = error "map_type";
-
-fun find_most_specific is_inst f eq xs T =
- find_first (fn U => is_inst (T, f U)
- andalso forall (fn U' => eq (f U, f U') orelse not
- (is_inst (T, f U') andalso is_inst (f U', f U)))
- xs) xs;
-
-fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of
- NONE => NONE
- | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T;
-
-fun lookup_rule thy f rules = find_most_specific
- (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules;
-
-fun infer_arities thy arities (optf, t) fs = case strip_comb t of
- (Abs (s, T, u), []) => infer_arities thy arities (NONE, u) fs
- | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs
- | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of
- SOME (SOME (_, (arity, _))) =>
- (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs
- handle Subscript => error "infer_arities: bad term")
- | _ => fold (infer_arities thy arities) (map (pair NONE) ts)
- (case optf of
- NONE => fs
- | SOME f => AList.update op = (u, the_default f
- (Option.map (curry op inter f) (AList.lookup op = fs u))) fs));
-
-
-(**************************************************************)
-(* derive the to_pred equation from the to_set equation *)
-(* *)
-(* 1. instantiate each set parameter with {(x, y). p x y} *)
-(* 2. apply %P. {(x, y). P x y} to both sides of the equation *)
-(* 3. simplify *)
-(**************************************************************)
-
-fun mk_to_pred_inst thy fs =
- map (fn (x, ps) =>
- let
- val U = HOLogic.dest_setT (fastype_of x);
- val x' = map_type (K (HOLogic.prodT_factors' ps U ---> HOLogic.boolT)) x
- in
- (cterm_of thy x,
- cterm_of thy (HOLogic.Collect_const U $
- HOLogic.ap_split' ps U HOLogic.boolT x'))
- end) fs;
-
-fun mk_to_pred_eq p fs optfs' T thm =
- let
- val thy = theory_of_thm thm;
- val insts = mk_to_pred_inst thy fs;
- val thm' = Thm.instantiate ([], insts) thm;
- val thm'' = (case optfs' of
- NONE => thm' RS sym
- | SOME fs' =>
- let
- val (_, U) = split_last (binder_types T);
- val Ts = HOLogic.prodT_factors' fs' U;
- (* FIXME: should cterm_instantiate increment indexes? *)
- val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
- val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
- Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
- in
- thm' RS (Drule.cterm_instantiate [(arg_cong_f,
- cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT,
- HOLogic.Collect_const U $ HOLogic.ap_split' fs' U
- HOLogic.boolT (Bound 0))))] arg_cong' RS sym)
- end)
- in
- Simplifier.simplify (HOL_basic_ss addsimps [mem_Collect_eq, split_conv]
- addsimprocs [collect_mem_simproc]) thm'' |>
- zero_var_indexes |> eta_contract_thm (equal p)
- end;
-
-
-(**** declare rules for converting predicates to sets ****)
-
-fun add ctxt thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
- case prop_of thm of
- Const ("Trueprop", _) $ (Const ("op =", Type (_, [T, _])) $ lhs $ rhs) =>
- (case body_type T of
- Type ("bool", []) =>
- let
- val thy = Context.theory_of ctxt;
- fun factors_of t fs = case strip_abs_body t of
- Const ("op :", _) $ u $ S =>
- if is_Free S orelse is_Var S then
- let val ps = HOLogic.prod_factors u
- in (SOME ps, (S, ps) :: fs) end
- else (NONE, fs)
- | _ => (NONE, fs);
- val (h, ts) = strip_comb lhs
- val (pfs, fs) = fold_map factors_of ts [];
- val ((h', ts'), fs') = (case rhs of
- Abs _ => (case strip_abs_body rhs of
- Const ("op :", _) $ u $ S =>
- (strip_comb S, SOME (HOLogic.prod_factors u))
- | _ => error "member symbol on right-hand side expected")
- | _ => (strip_comb rhs, NONE))
- in
- case (name_type_of h, name_type_of h') of
- (SOME (s, T), SOME (s', T')) =>
- if exists (fn (U, _) =>
- Sign.typ_instance thy (T', U) andalso
- Sign.typ_instance thy (U, T'))
- (Symtab.lookup_list set_arities s')
- then
- (warning ("Ignoring conversion rule for operator " ^ s'); tab)
- else
- {to_set_simps = thm :: to_set_simps,
- to_pred_simps =
- mk_to_pred_eq h fs fs' T' thm :: to_pred_simps,
- set_arities = Symtab.insert_list op = (s',
- (T', (map (AList.lookup op = fs) ts', fs'))) set_arities,
- pred_arities = Symtab.insert_list op = (s,
- (T, (pfs, fs'))) pred_arities}
- | _ => error "set / predicate constant expected"
- end
- | _ => error "equation between predicates expected")
- | _ => error "equation expected";
-
-val pred_set_conv_att = Thm.declaration_attribute
- (fn thm => fn ctxt => PredSetConvData.map (add ctxt thm) ctxt);
-
-
-(**** convert theorem in set notation to predicate notation ****)
-
-fun is_pred tab t =
- case Option.map (Symtab.lookup tab o fst) (name_type_of t) of
- SOME (SOME _) => true | _ => false;
-
-fun to_pred_simproc rules =
- let val rules' = map mk_meta_eq rules
- in
- Simplifier.simproc_i @{theory HOL} "to_pred" [anyt]
- (fn thy => K (lookup_rule thy (prop_of #> Logic.dest_equals) rules'))
- end;
-
-fun to_pred_proc thy rules t = case lookup_rule thy I rules t of
- NONE => NONE
- | SOME (lhs, rhs) =>
- SOME (Envir.subst_vars
- (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs);
-
-fun to_pred thms ctxt thm =
- let
- val thy = Context.theory_of ctxt;
- val {to_pred_simps, set_arities, pred_arities, ...} =
- fold (add ctxt) thms (PredSetConvData.get ctxt);
- val fs = filter (is_Var o fst)
- (infer_arities thy set_arities (NONE, prop_of thm) []);
- (* instantiate each set parameter with {(x, y). p x y} *)
- val insts = mk_to_pred_inst thy fs
- in
- thm |>
- Thm.instantiate ([], insts) |>
- Simplifier.full_simplify (HOL_basic_ss addsimprocs
- [to_pred_simproc (mem_Collect_eq :: split_conv :: to_pred_simps)]) |>
- eta_contract_thm (is_pred pred_arities) |>
- RuleCases.save thm
- end;
-
-val to_pred_att = Thm.rule_attribute o to_pred;
-
-
-(**** convert theorem in predicate notation to set notation ****)
-
-fun to_set thms ctxt thm =
- let
- val thy = Context.theory_of ctxt;
- val {to_set_simps, pred_arities, ...} =
- fold (add ctxt) thms (PredSetConvData.get ctxt);
- val fs = filter (is_Var o fst)
- (infer_arities thy pred_arities (NONE, prop_of thm) []);
- (* instantiate each predicate parameter with %x y. (x, y) : s *)
- val insts = map (fn (x, ps) =>
- let
- val Ts = binder_types (fastype_of x);
- val T = HOLogic.mk_tupleT ps Ts;
- val x' = map_type (K (HOLogic.mk_setT T)) x
- in
- (cterm_of thy x,
- cterm_of thy (list_abs (map (pair "x") Ts, HOLogic.mk_mem
- (HOLogic.mk_tuple' ps T (map Bound (length ps downto 0)), x'))))
- end) fs
- in
- thm |>
- Thm.instantiate ([], insts) |>
- Simplifier.full_simplify (HOL_basic_ss addsimps to_set_simps
- addsimprocs [strong_ind_simproc pred_arities, collect_mem_simproc]) |>
- RuleCases.save thm
- end;
-
-val to_set_att = Thm.rule_attribute o to_set;
-
-
-(**** preprocessor for code generator ****)
-
-fun codegen_preproc thy =
- let
- val {to_pred_simps, set_arities, pred_arities, ...} =
- PredSetConvData.get (Context.Theory thy);
- fun preproc thm =
- if exists_Const (fn (s, _) => case Symtab.lookup set_arities s of
- NONE => false
- | SOME arities => exists (fn (_, (xs, _)) =>
- forall is_none xs) arities) (prop_of thm)
- then
- thm |>
- Simplifier.full_simplify (HOL_basic_ss addsimprocs
- [to_pred_simproc (mem_Collect_eq :: split_conv :: to_pred_simps)]) |>
- eta_contract_thm (is_pred pred_arities)
- else thm
- in map preproc end;
-
-fun code_ind_att optmod = to_pred_att [] #> InductiveCodegen.add optmod NONE;
-
-
-(**** definition of inductive sets ****)
-
-fun add_ind_set_def
- {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
- cs intros monos params cnames_syn ctxt =
- let
- val thy = ProofContext.theory_of ctxt;
- val {set_arities, pred_arities, to_pred_simps, ...} =
- PredSetConvData.get (Context.Proof ctxt);
- fun infer (Abs (_, _, t)) = infer t
- | infer (Const ("op :", _) $ t $ u) =
- infer_arities thy set_arities (SOME (HOLogic.prod_factors t), u)
- | infer (t $ u) = infer t #> infer u
- | infer _ = I;
- val new_arities = filter_out
- (fn (x as Free (_, T), _) => x mem params andalso length (binder_types T) > 1
- | _ => false) (fold (snd #> infer) intros []);
- val params' = map (fn x => (case AList.lookup op = new_arities x of
- SOME fs =>
- let
- val T = HOLogic.dest_setT (fastype_of x);
- val Ts = HOLogic.prodT_factors' fs T;
- val x' = map_type (K (Ts ---> HOLogic.boolT)) x
- in
- (x, (x',
- (HOLogic.Collect_const T $
- HOLogic.ap_split' fs T HOLogic.boolT x',
- list_abs (map (pair "x") Ts, HOLogic.mk_mem
- (HOLogic.mk_tuple' fs T (map Bound (length fs downto 0)),
- x)))))
- end
- | NONE => (x, (x, (x, x))))) params;
- val (params1, (params2, params3)) =
- params' |> map snd |> split_list ||> split_list;
- val paramTs = map fastype_of params;
-
- (* equations for converting sets to predicates *)
- val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) =>
- let
- val fs = the_default [] (AList.lookup op = new_arities c);
- val (Us, U) = split_last (binder_types T);
- val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks
- [Pretty.str "Argument types",
- Pretty.block (Pretty.commas (map (Syntax.pretty_typ ctxt) Us)),
- Pretty.str ("of " ^ s ^ " do not agree with types"),
- Pretty.block (Pretty.commas (map (Syntax.pretty_typ ctxt) paramTs)),
- Pretty.str "of declared parameters"]));
- val Ts = HOLogic.prodT_factors' fs U;
- val c' = Free (s ^ "p",
- map fastype_of params1 @ Ts ---> HOLogic.boolT)
- in
- ((c', (fs, U, Ts)),
- (list_comb (c, params2),
- HOLogic.Collect_const U $ HOLogic.ap_split' fs U HOLogic.boolT
- (list_comb (c', params1))))
- end) |> split_list |>> split_list;
- val eqns' = eqns @
- map (prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq)
- (mem_Collect_eq :: split_conv :: to_pred_simps);
-
- (* predicate version of the introduction rules *)
- val intros' =
- map (fn (name_atts, t) => (name_atts,
- t |>
- map_aterms (fn u =>
- (case AList.lookup op = params' u of
- SOME (_, (u', _)) => u'
- | NONE => u)) |>
- Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |>
- eta_contract (member op = cs' orf is_pred pred_arities))) intros;
- val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn;
- val monos' = map (to_pred [] (Context.Proof ctxt)) monos;
- val ({preds, intrs, elims, raw_induct, ...}, ctxt1) =
- InductivePackage.add_ind_def
- {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Binding.empty,
- coind = coind, no_elim = no_elim, no_ind = no_ind,
- skip_mono = skip_mono, fork_mono = fork_mono}
- cs' intros' monos' params1 cnames_syn' ctxt;
-
- (* define inductive sets using previously defined predicates *)
- val (defs, ctxt2) = fold_map (LocalTheory.define Thm.internalK)
- (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.empty_binding,
- fold_rev lambda params (HOLogic.Collect_const U $
- HOLogic.ap_split' fs U HOLogic.boolT (list_comb (p, params3))))))
- (cnames_syn ~~ cs_info ~~ preds)) ctxt1;
-
- (* prove theorems for converting predicate to set notation *)
- val ctxt3 = fold
- (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn ctxt =>
- let val conv_thm =
- Goal.prove ctxt (map (fst o dest_Free) params) []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (list_comb (p, params3),
- list_abs (map (pair "x") Ts, HOLogic.mk_mem
- (HOLogic.mk_tuple' fs U (map Bound (length fs downto 0)),
- list_comb (c, params))))))
- (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps
- [def, mem_Collect_eq, split_conv]) 1))
- in
- ctxt |> LocalTheory.note kind ((Binding.name (s ^ "p_" ^ s ^ "_eq"),
- [Attrib.internal (K pred_set_conv_att)]),
- [conv_thm]) |> snd
- end) (preds ~~ cs ~~ cs_info ~~ defs) ctxt2;
-
- (* convert theorems to set notation *)
- val rec_name =
- if Binding.is_empty alt_name then
- Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
- else alt_name;
- val cnames = map (LocalTheory.full_name ctxt3 o #1) cnames_syn; (* FIXME *)
- val (intr_names, intr_atts) = split_list (map fst intros);
- val raw_induct' = to_set [] (Context.Proof ctxt3) raw_induct;
- val (intrs', elims', induct, ctxt4) =
- InductivePackage.declare_rules kind rec_name coind no_ind cnames
- (map (to_set [] (Context.Proof ctxt3)) intrs) intr_names intr_atts
- (map (fn th => (to_set [] (Context.Proof ctxt3) th,
- map fst (fst (RuleCases.get th)))) elims)
- raw_induct' ctxt3
- in
- ({intrs = intrs', elims = elims', induct = induct,
- raw_induct = raw_induct', preds = map fst defs},
- ctxt4)
- end;
-
-val add_inductive_i = InductivePackage.gen_add_inductive_i add_ind_set_def;
-val add_inductive = InductivePackage.gen_add_inductive add_ind_set_def;
-
-val mono_add_att = to_pred_att [] #> InductivePackage.mono_add;
-val mono_del_att = to_pred_att [] #> InductivePackage.mono_del;
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- Attrib.setup @{binding pred_set_conv} (Scan.succeed pred_set_conv_att)
- "declare rules for converting between predicate and set notation" #>
- Attrib.setup @{binding to_set} (Attrib.thms >> to_set_att) "convert rule to set notation" #>
- Attrib.setup @{binding to_pred} (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #>
- Code.add_attribute ("ind_set",
- Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> code_ind_att) #>
- Codegen.add_preprocessor codegen_preproc #>
- Attrib.setup @{binding mono_set} (Attrib.add_del mono_add_att mono_del_att)
- "declaration of monotonicity rule for set operators" #>
- Context.theory_map (Simplifier.map_ss (fn ss => ss addsimprocs [collect_mem_simproc]));
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val ind_set_decl = InductivePackage.gen_ind_decl add_ind_set_def;
-
-val _ =
- OuterSyntax.local_theory' "inductive_set" "define inductive sets" K.thy_decl (ind_set_decl false);
-
-val _ =
- OuterSyntax.local_theory' "coinductive_set" "define coinductive sets" K.thy_decl (ind_set_decl true);
-
-end;
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/old_primrec.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,348 @@
+(* Title: HOL/Tools/old_primrec.ML
+ Author: Norbert Voelker, FernUni Hagen
+ Author: Stefan Berghofer, TU Muenchen
+
+Package for defining functions on datatypes by primitive recursion.
+*)
+
+signature OLD_PRIMREC =
+sig
+ val unify_consts: theory -> term list -> term list -> term list * term list
+ val add_primrec: string -> ((bstring * string) * Attrib.src list) list
+ -> theory -> thm list * theory
+ val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list
+ -> theory -> thm list * theory
+ val add_primrec_i: string -> ((bstring * term) * attribute list) list
+ -> theory -> thm list * theory
+ val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list
+ -> theory -> thm list * theory
+end;
+
+structure OldPrimrec : OLD_PRIMREC =
+struct
+
+open DatatypeAux;
+
+exception RecError of string;
+
+fun primrec_err s = error ("Primrec definition error:\n" ^ s);
+fun primrec_eq_err thy s eq =
+ primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq));
+
+
+(*the following code ensures that each recursive set always has the
+ same type in all introduction rules*)
+fun unify_consts thy cs intr_ts =
+ (let
+ fun varify (t, (i, ts)) =
+ let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t))
+ in (maxidx_of_term t', t'::ts) end;
+ val (i, cs') = List.foldr varify (~1, []) cs;
+ val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
+ val rec_consts = fold Term.add_consts cs' [];
+ val intr_consts = fold Term.add_consts intr_ts' [];
+ fun unify (cname, cT) =
+ let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts)
+ in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
+ val (env, _) = fold unify rec_consts (Vartab.empty, i');
+ val subst = Type.freeze o map_types (Envir.norm_type env)
+
+ in (map subst cs', map subst intr_ts')
+ end) handle Type.TUNIFY =>
+ (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
+
+
+(* preprocessing of equations *)
+
+fun process_eqn thy eq rec_fns =
+ let
+ val (lhs, rhs) =
+ if null (Term.add_vars eq []) then
+ HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
+ handle TERM _ => raise RecError "not a proper equation"
+ else raise RecError "illegal schematic variable(s)";
+
+ val (recfun, args) = strip_comb lhs;
+ val fnameT = dest_Const recfun handle TERM _ =>
+ raise RecError "function is not declared as constant in theory";
+
+ val (ls', rest) = take_prefix is_Free args;
+ val (middle, rs') = take_suffix is_Free rest;
+ val rpos = length ls';
+
+ val (constr, cargs') = if null middle then raise RecError "constructor missing"
+ else strip_comb (hd middle);
+ val (cname, T) = dest_Const constr
+ handle TERM _ => raise RecError "ill-formed constructor";
+ val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
+ raise RecError "cannot determine datatype associated with function"
+
+ val (ls, cargs, rs) =
+ (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
+ handle TERM _ => raise RecError "illegal argument in pattern";
+ val lfrees = ls @ rs @ cargs;
+
+ fun check_vars _ [] = ()
+ | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars))
+ in
+ if length middle > 1 then
+ raise RecError "more than one non-variable in pattern"
+ else
+ (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
+ check_vars "extra variables on rhs: "
+ (map dest_Free (OldTerm.term_frees rhs) \\ lfrees);
+ case AList.lookup (op =) rec_fns fnameT of
+ NONE =>
+ (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
+ | SOME (_, rpos', eqns) =>
+ if AList.defined (op =) eqns cname then
+ raise RecError "constructor already occurred as pattern"
+ else if rpos <> rpos' then
+ raise RecError "position of recursive argument inconsistent"
+ else
+ AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns))
+ rec_fns)
+ end
+ handle RecError s => primrec_eq_err thy s eq;
+
+fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) =
+ let
+ val (_, (tname, _, constrs)) = List.nth (descr, i);
+
+ (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
+
+ fun subst [] t fs = (t, fs)
+ | subst subs (Abs (a, T, t)) fs =
+ fs
+ |> subst subs t
+ |-> (fn t' => pair (Abs (a, T, t')))
+ | subst subs (t as (_ $ _)) fs =
+ let
+ val (f, ts) = strip_comb t;
+ in
+ if is_Const f andalso dest_Const f mem map fst rec_eqns then
+ let
+ val fnameT' as (fname', _) = dest_Const f;
+ val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT');
+ val ls = Library.take (rpos, ts);
+ val rest = Library.drop (rpos, ts);
+ val (x', rs) = (hd rest, tl rest)
+ handle Empty => raise RecError ("not enough arguments\
+ \ in recursive application\nof function " ^ quote fname' ^ " on rhs");
+ val (x, xs) = strip_comb x'
+ in case AList.lookup (op =) subs x
+ of NONE =>
+ fs
+ |> fold_map (subst subs) ts
+ |-> (fn ts' => pair (list_comb (f, ts')))
+ | SOME (i', y) =>
+ fs
+ |> fold_map (subst subs) (xs @ ls @ rs)
+ ||> process_fun thy descr rec_eqns (i', fnameT')
+ |-> (fn ts' => pair (list_comb (y, ts')))
+ end
+ else
+ fs
+ |> fold_map (subst subs) (f :: ts)
+ |-> (fn (f'::ts') => pair (list_comb (f', ts')))
+ end
+ | subst _ t fs = (t, fs);
+
+ (* translate rec equations into function arguments suitable for rec comb *)
+
+ fun trans eqns (cname, cargs) (fnameTs', fnss', fns) =
+ (case AList.lookup (op =) eqns cname of
+ NONE => (warning ("No equation for constructor " ^ quote cname ^
+ "\nin definition of function " ^ quote fname);
+ (fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns))
+ | SOME (ls, cargs', rs, rhs, eq) =>
+ let
+ val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
+ val rargs = map fst recs;
+ val subs = map (rpair dummyT o fst)
+ (rev (Term.rename_wrt_term rhs rargs));
+ val (rhs', (fnameTs'', fnss'')) =
+ (subst (map (fn ((x, y), z) =>
+ (Free x, (body_index y, Free z)))
+ (recs ~~ subs)) rhs (fnameTs', fnss'))
+ handle RecError s => primrec_eq_err thy s eq
+ in (fnameTs'', fnss'',
+ (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns)
+ end)
+
+ in (case AList.lookup (op =) fnameTs i of
+ NONE =>
+ if exists (equal fnameT o snd) fnameTs then
+ raise RecError ("inconsistent functions for datatype " ^ quote tname)
+ else
+ let
+ val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT);
+ val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs
+ ((i, fnameT)::fnameTs, fnss, [])
+ in
+ (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
+ end
+ | SOME fnameT' =>
+ if fnameT = fnameT' then (fnameTs, fnss)
+ else raise RecError ("inconsistent functions for datatype " ^ quote tname))
+ end;
+
+
+(* prepare functions needed for definitions *)
+
+fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
+ case AList.lookup (op =) fns i of
+ NONE =>
+ let
+ val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
+ replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
+ dummyT ---> HOLogic.unitT)) constrs;
+ val _ = warning ("No function definition for datatype " ^ quote tname)
+ in
+ (dummy_fns @ fs, defs)
+ end
+ | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs);
+
+
+(* make definition *)
+
+fun make_def thy fs (fname, ls, rec_name, tname) =
+ let
+ val rhs = fold_rev (fn T => fn t => Abs ("", T, t))
+ ((map snd ls) @ [dummyT])
+ (list_comb (Const (rec_name, dummyT),
+ fs @ map Bound (0 ::(length ls downto 1))))
+ val def_name = Long_Name.base_name fname ^ "_" ^ Long_Name.base_name tname ^ "_def";
+ val def_prop =
+ singleton (Syntax.check_terms (ProofContext.init thy))
+ (Logic.mk_equals (Const (fname, dummyT), rhs));
+ in (def_name, def_prop) end;
+
+
+(* find datatypes which contain all datatypes in tnames' *)
+
+fun find_dts (dt_info : info Symtab.table) _ [] = []
+ | find_dts dt_info tnames' (tname::tnames) =
+ (case Symtab.lookup dt_info tname of
+ NONE => primrec_err (quote tname ^ " is not a datatype")
+ | SOME dt =>
+ if tnames' subset (map (#1 o snd) (#descr dt)) then
+ (tname, dt)::(find_dts dt_info tnames' tnames)
+ else find_dts dt_info tnames' tnames);
+
+fun prepare_induct ({descr, induction, ...}: info) rec_eqns =
+ let
+ fun constrs_of (_, (_, _, cs)) =
+ map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs;
+ val params_of = these o AList.lookup (op =) (List.concat (map constrs_of rec_eqns));
+ in
+ induction
+ |> RuleCases.rename_params (map params_of (List.concat (map (map #1 o #3 o #2) descr)))
+ |> RuleCases.save induction
+ end;
+
+local
+
+fun gen_primrec_i note def alt_name eqns_atts thy =
+ let
+ val (eqns, atts) = split_list eqns_atts;
+ val dt_info = Datatype.get_datatypes thy;
+ val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ;
+ val tnames = distinct (op =) (map (#1 o snd) rec_eqns);
+ val dts = find_dts dt_info tnames tnames;
+ val main_fns =
+ map (fn (tname, {index, ...}) =>
+ (index,
+ (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns))
+ dts;
+ val {descr, rec_names, rec_rewrites, ...} =
+ if null dts then
+ primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
+ else snd (hd dts);
+ val (fnameTs, fnss) =
+ fold_rev (process_fun thy descr rec_eqns) main_fns ([], []);
+ val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
+ val defs' = map (make_def thy fs) defs;
+ val nameTs1 = map snd fnameTs;
+ val nameTs2 = map fst rec_eqns;
+ val _ = if gen_eq_set (op =) (nameTs1, nameTs2) then ()
+ else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^
+ "\nare not mutually recursive");
+ val primrec_name =
+ if alt_name = "" then (space_implode "_" (map (Long_Name.base_name o #1) defs)) else alt_name;
+ val (defs_thms', thy') =
+ thy
+ |> Sign.add_path primrec_name
+ |> fold_map def (map (fn (name, t) => ((name, []), t)) defs');
+ val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms';
+ val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t
+ (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
+ val (simps', thy'') =
+ thy'
+ |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps);
+ val simps'' = maps snd simps';
+ in
+ thy''
+ |> note (("simps", [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
+ Code.add_default_eqn_attribute]), simps'')
+ |> snd
+ |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns])
+ |> snd
+ |> Sign.parent_path
+ |> pair simps''
+ end;
+
+fun gen_primrec note def alt_name eqns thy =
+ let
+ val ((names, strings), srcss) = apfst split_list (split_list eqns);
+ val atts = map (map (Attrib.attribute thy)) srcss;
+ val eqn_ts = map (fn s => Syntax.read_prop_global thy s
+ handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings;
+ val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq)))
+ handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts;
+ val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts
+ in
+ gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy
+ end;
+
+fun thy_note ((name, atts), thms) =
+ PureThy.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms));
+fun thy_def false ((name, atts), t) =
+ PureThy.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm))
+ | thy_def true ((name, atts), t) =
+ PureThy.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm));
+
+in
+
+val add_primrec = gen_primrec thy_note (thy_def false);
+val add_primrec_unchecked = gen_primrec thy_note (thy_def true);
+val add_primrec_i = gen_primrec_i thy_note (thy_def false);
+val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true);
+fun gen_primrec note def alt_name specs =
+ gen_primrec_i note def alt_name (map (fn ((name, t), atts) => ((name, atts), t)) specs);
+
+end;
+
+
+(* see primrec.ML (* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val opt_unchecked_name =
+ Scan.optional (P.$$$ "(" |-- P.!!!
+ (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
+ P.name >> pair false) --| P.$$$ ")")) (false, "");
+
+val primrec_decl =
+ opt_unchecked_name -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop);
+
+val _ =
+ OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
+ (primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
+ Toplevel.theory (snd o
+ (if unchecked then add_primrec_unchecked else add_primrec) alt_name
+ (map P.triple_swap eqns))));
+
+end;*)
+
+end;
--- a/src/HOL/Tools/old_primrec_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,348 +0,0 @@
-(* Title: HOL/Tools/old_primrec_package.ML
- Author: Norbert Voelker, FernUni Hagen
- Author: Stefan Berghofer, TU Muenchen
-
-Package for defining functions on datatypes by primitive recursion.
-*)
-
-signature OLD_PRIMREC_PACKAGE =
-sig
- val unify_consts: theory -> term list -> term list -> term list * term list
- val add_primrec: string -> ((bstring * string) * Attrib.src list) list
- -> theory -> thm list * theory
- val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list
- -> theory -> thm list * theory
- val add_primrec_i: string -> ((bstring * term) * attribute list) list
- -> theory -> thm list * theory
- val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list
- -> theory -> thm list * theory
-end;
-
-structure OldPrimrecPackage : OLD_PRIMREC_PACKAGE =
-struct
-
-open DatatypeAux;
-
-exception RecError of string;
-
-fun primrec_err s = error ("Primrec definition error:\n" ^ s);
-fun primrec_eq_err thy s eq =
- primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq));
-
-
-(*the following code ensures that each recursive set always has the
- same type in all introduction rules*)
-fun unify_consts thy cs intr_ts =
- (let
- fun varify (t, (i, ts)) =
- let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t))
- in (maxidx_of_term t', t'::ts) end;
- val (i, cs') = List.foldr varify (~1, []) cs;
- val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
- val rec_consts = fold Term.add_consts cs' [];
- val intr_consts = fold Term.add_consts intr_ts' [];
- fun unify (cname, cT) =
- let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts)
- in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
- val (env, _) = fold unify rec_consts (Vartab.empty, i');
- val subst = Type.freeze o map_types (Envir.norm_type env)
-
- in (map subst cs', map subst intr_ts')
- end) handle Type.TUNIFY =>
- (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
-
-
-(* preprocessing of equations *)
-
-fun process_eqn thy eq rec_fns =
- let
- val (lhs, rhs) =
- if null (Term.add_vars eq []) then
- HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
- handle TERM _ => raise RecError "not a proper equation"
- else raise RecError "illegal schematic variable(s)";
-
- val (recfun, args) = strip_comb lhs;
- val fnameT = dest_Const recfun handle TERM _ =>
- raise RecError "function is not declared as constant in theory";
-
- val (ls', rest) = take_prefix is_Free args;
- val (middle, rs') = take_suffix is_Free rest;
- val rpos = length ls';
-
- val (constr, cargs') = if null middle then raise RecError "constructor missing"
- else strip_comb (hd middle);
- val (cname, T) = dest_Const constr
- handle TERM _ => raise RecError "ill-formed constructor";
- val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
- raise RecError "cannot determine datatype associated with function"
-
- val (ls, cargs, rs) =
- (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
- handle TERM _ => raise RecError "illegal argument in pattern";
- val lfrees = ls @ rs @ cargs;
-
- fun check_vars _ [] = ()
- | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars))
- in
- if length middle > 1 then
- raise RecError "more than one non-variable in pattern"
- else
- (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
- check_vars "extra variables on rhs: "
- (map dest_Free (OldTerm.term_frees rhs) \\ lfrees);
- case AList.lookup (op =) rec_fns fnameT of
- NONE =>
- (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
- | SOME (_, rpos', eqns) =>
- if AList.defined (op =) eqns cname then
- raise RecError "constructor already occurred as pattern"
- else if rpos <> rpos' then
- raise RecError "position of recursive argument inconsistent"
- else
- AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns))
- rec_fns)
- end
- handle RecError s => primrec_eq_err thy s eq;
-
-fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) =
- let
- val (_, (tname, _, constrs)) = List.nth (descr, i);
-
- (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
-
- fun subst [] t fs = (t, fs)
- | subst subs (Abs (a, T, t)) fs =
- fs
- |> subst subs t
- |-> (fn t' => pair (Abs (a, T, t')))
- | subst subs (t as (_ $ _)) fs =
- let
- val (f, ts) = strip_comb t;
- in
- if is_Const f andalso dest_Const f mem map fst rec_eqns then
- let
- val fnameT' as (fname', _) = dest_Const f;
- val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT');
- val ls = Library.take (rpos, ts);
- val rest = Library.drop (rpos, ts);
- val (x', rs) = (hd rest, tl rest)
- handle Empty => raise RecError ("not enough arguments\
- \ in recursive application\nof function " ^ quote fname' ^ " on rhs");
- val (x, xs) = strip_comb x'
- in case AList.lookup (op =) subs x
- of NONE =>
- fs
- |> fold_map (subst subs) ts
- |-> (fn ts' => pair (list_comb (f, ts')))
- | SOME (i', y) =>
- fs
- |> fold_map (subst subs) (xs @ ls @ rs)
- ||> process_fun thy descr rec_eqns (i', fnameT')
- |-> (fn ts' => pair (list_comb (y, ts')))
- end
- else
- fs
- |> fold_map (subst subs) (f :: ts)
- |-> (fn (f'::ts') => pair (list_comb (f', ts')))
- end
- | subst _ t fs = (t, fs);
-
- (* translate rec equations into function arguments suitable for rec comb *)
-
- fun trans eqns (cname, cargs) (fnameTs', fnss', fns) =
- (case AList.lookup (op =) eqns cname of
- NONE => (warning ("No equation for constructor " ^ quote cname ^
- "\nin definition of function " ^ quote fname);
- (fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns))
- | SOME (ls, cargs', rs, rhs, eq) =>
- let
- val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
- val rargs = map fst recs;
- val subs = map (rpair dummyT o fst)
- (rev (Term.rename_wrt_term rhs rargs));
- val (rhs', (fnameTs'', fnss'')) =
- (subst (map (fn ((x, y), z) =>
- (Free x, (body_index y, Free z)))
- (recs ~~ subs)) rhs (fnameTs', fnss'))
- handle RecError s => primrec_eq_err thy s eq
- in (fnameTs'', fnss'',
- (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns)
- end)
-
- in (case AList.lookup (op =) fnameTs i of
- NONE =>
- if exists (equal fnameT o snd) fnameTs then
- raise RecError ("inconsistent functions for datatype " ^ quote tname)
- else
- let
- val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT);
- val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs
- ((i, fnameT)::fnameTs, fnss, [])
- in
- (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
- end
- | SOME fnameT' =>
- if fnameT = fnameT' then (fnameTs, fnss)
- else raise RecError ("inconsistent functions for datatype " ^ quote tname))
- end;
-
-
-(* prepare functions needed for definitions *)
-
-fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
- case AList.lookup (op =) fns i of
- NONE =>
- let
- val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
- replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
- dummyT ---> HOLogic.unitT)) constrs;
- val _ = warning ("No function definition for datatype " ^ quote tname)
- in
- (dummy_fns @ fs, defs)
- end
- | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs);
-
-
-(* make definition *)
-
-fun make_def thy fs (fname, ls, rec_name, tname) =
- let
- val rhs = fold_rev (fn T => fn t => Abs ("", T, t))
- ((map snd ls) @ [dummyT])
- (list_comb (Const (rec_name, dummyT),
- fs @ map Bound (0 ::(length ls downto 1))))
- val def_name = Long_Name.base_name fname ^ "_" ^ Long_Name.base_name tname ^ "_def";
- val def_prop =
- singleton (Syntax.check_terms (ProofContext.init thy))
- (Logic.mk_equals (Const (fname, dummyT), rhs));
- in (def_name, def_prop) end;
-
-
-(* find datatypes which contain all datatypes in tnames' *)
-
-fun find_dts (dt_info : datatype_info Symtab.table) _ [] = []
- | find_dts dt_info tnames' (tname::tnames) =
- (case Symtab.lookup dt_info tname of
- NONE => primrec_err (quote tname ^ " is not a datatype")
- | SOME dt =>
- if tnames' subset (map (#1 o snd) (#descr dt)) then
- (tname, dt)::(find_dts dt_info tnames' tnames)
- else find_dts dt_info tnames' tnames);
-
-fun prepare_induct ({descr, induction, ...}: datatype_info) rec_eqns =
- let
- fun constrs_of (_, (_, _, cs)) =
- map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs;
- val params_of = these o AList.lookup (op =) (List.concat (map constrs_of rec_eqns));
- in
- induction
- |> RuleCases.rename_params (map params_of (List.concat (map (map #1 o #3 o #2) descr)))
- |> RuleCases.save induction
- end;
-
-local
-
-fun gen_primrec_i note def alt_name eqns_atts thy =
- let
- val (eqns, atts) = split_list eqns_atts;
- val dt_info = DatatypePackage.get_datatypes thy;
- val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ;
- val tnames = distinct (op =) (map (#1 o snd) rec_eqns);
- val dts = find_dts dt_info tnames tnames;
- val main_fns =
- map (fn (tname, {index, ...}) =>
- (index,
- (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns))
- dts;
- val {descr, rec_names, rec_rewrites, ...} =
- if null dts then
- primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
- else snd (hd dts);
- val (fnameTs, fnss) =
- fold_rev (process_fun thy descr rec_eqns) main_fns ([], []);
- val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
- val defs' = map (make_def thy fs) defs;
- val nameTs1 = map snd fnameTs;
- val nameTs2 = map fst rec_eqns;
- val _ = if gen_eq_set (op =) (nameTs1, nameTs2) then ()
- else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^
- "\nare not mutually recursive");
- val primrec_name =
- if alt_name = "" then (space_implode "_" (map (Long_Name.base_name o #1) defs)) else alt_name;
- val (defs_thms', thy') =
- thy
- |> Sign.add_path primrec_name
- |> fold_map def (map (fn (name, t) => ((name, []), t)) defs');
- val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms';
- val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t
- (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
- val (simps', thy'') =
- thy'
- |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps);
- val simps'' = maps snd simps';
- in
- thy''
- |> note (("simps", [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
- Code.add_default_eqn_attribute]), simps'')
- |> snd
- |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns])
- |> snd
- |> Sign.parent_path
- |> pair simps''
- end;
-
-fun gen_primrec note def alt_name eqns thy =
- let
- val ((names, strings), srcss) = apfst split_list (split_list eqns);
- val atts = map (map (Attrib.attribute thy)) srcss;
- val eqn_ts = map (fn s => Syntax.read_prop_global thy s
- handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings;
- val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq)))
- handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts;
- val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts
- in
- gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy
- end;
-
-fun thy_note ((name, atts), thms) =
- PureThy.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms));
-fun thy_def false ((name, atts), t) =
- PureThy.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm))
- | thy_def true ((name, atts), t) =
- PureThy.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm));
-
-in
-
-val add_primrec = gen_primrec thy_note (thy_def false);
-val add_primrec_unchecked = gen_primrec thy_note (thy_def true);
-val add_primrec_i = gen_primrec_i thy_note (thy_def false);
-val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true);
-fun gen_primrec note def alt_name specs =
- gen_primrec_i note def alt_name (map (fn ((name, t), atts) => ((name, atts), t)) specs);
-
-end;
-
-
-(* see primrecr_package.ML (* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_unchecked_name =
- Scan.optional (P.$$$ "(" |-- P.!!!
- (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
- P.name >> pair false) --| P.$$$ ")")) (false, "");
-
-val primrec_decl =
- opt_unchecked_name -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop);
-
-val _ =
- OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
- (primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
- Toplevel.theory (snd o
- (if unchecked then add_primrec_unchecked else add_primrec) alt_name
- (map P.triple_swap eqns))));
-
-end;*)
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/primrec.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,332 @@
+(* Title: HOL/Tools/primrec.ML
+ Author: Stefan Berghofer, TU Muenchen; Norbert Voelker, FernUni Hagen;
+ Florian Haftmann, TU Muenchen
+
+Package for defining functions on datatypes by primitive recursion.
+*)
+
+signature PRIMREC =
+sig
+ val add_primrec: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> local_theory -> thm list * local_theory
+ val add_primrec_cmd: (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> local_theory -> thm list * local_theory
+ val add_primrec_global: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> thm list * theory
+ val add_primrec_overloaded: (string * (string * typ) * bool) list ->
+ (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> thm list * theory
+ val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+ local_theory -> (string * thm list list) * local_theory
+end;
+
+structure Primrec : PRIMREC =
+struct
+
+open DatatypeAux;
+
+exception PrimrecError of string * term option;
+
+fun primrec_error msg = raise PrimrecError (msg, NONE);
+fun primrec_error_eqn msg eqn = raise PrimrecError (msg, SOME eqn);
+
+fun message s = if ! Toplevel.debug then tracing s else ();
+
+
+(* preprocessing of equations *)
+
+fun process_eqn is_fixed spec rec_fns =
+ let
+ val (vs, Ts) = split_list (strip_qnt_vars "all" spec);
+ val body = strip_qnt_body "all" spec;
+ val (vs', _) = Name.variants vs (Name.make_context (fold_aterms
+ (fn Free (v, _) => insert (op =) v | _ => I) body []));
+ val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body;
+ val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn)
+ handle TERM _ => primrec_error "not a proper equation";
+ val (recfun, args) = strip_comb lhs;
+ val fname = case recfun of Free (v, _) => if is_fixed v then v
+ else primrec_error "illegal head of function equation"
+ | _ => primrec_error "illegal head of function equation";
+
+ val (ls', rest) = take_prefix is_Free args;
+ val (middle, rs') = take_suffix is_Free rest;
+ val rpos = length ls';
+
+ val (constr, cargs') = if null middle then primrec_error "constructor missing"
+ else strip_comb (hd middle);
+ val (cname, T) = dest_Const constr
+ handle TERM _ => primrec_error "ill-formed constructor";
+ val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
+ primrec_error "cannot determine datatype associated with function"
+
+ val (ls, cargs, rs) =
+ (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
+ handle TERM _ => primrec_error "illegal argument in pattern";
+ val lfrees = ls @ rs @ cargs;
+
+ fun check_vars _ [] = ()
+ | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn;
+ in
+ if length middle > 1 then
+ primrec_error "more than one non-variable in pattern"
+ else
+ (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
+ check_vars "extra variables on rhs: "
+ (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
+ |> filter_out (is_fixed o fst));
+ case AList.lookup (op =) rec_fns fname of
+ NONE =>
+ (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))]))::rec_fns
+ | SOME (_, rpos', eqns) =>
+ if AList.defined (op =) eqns cname then
+ primrec_error "constructor already occurred as pattern"
+ else if rpos <> rpos' then
+ primrec_error "position of recursive argument inconsistent"
+ else
+ AList.update (op =)
+ (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn))::eqns))
+ rec_fns)
+ end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec;
+
+fun process_fun descr eqns (i, fname) (fnames, fnss) =
+ let
+ val (_, (tname, _, constrs)) = nth descr i;
+
+ (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
+
+ fun subst [] t fs = (t, fs)
+ | subst subs (Abs (a, T, t)) fs =
+ fs
+ |> subst subs t
+ |-> (fn t' => pair (Abs (a, T, t')))
+ | subst subs (t as (_ $ _)) fs =
+ let
+ val (f, ts) = strip_comb t;
+ in
+ if is_Free f
+ andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then
+ let
+ val (fname', _) = dest_Free f;
+ val (_, rpos, _) = the (AList.lookup (op =) eqns fname');
+ val (ls, rs) = chop rpos ts
+ val (x', rs') = case rs
+ of x' :: rs => (x', rs)
+ | [] => primrec_error ("not enough arguments in recursive application\n"
+ ^ "of function " ^ quote fname' ^ " on rhs");
+ val (x, xs) = strip_comb x';
+ in case AList.lookup (op =) subs x
+ of NONE =>
+ fs
+ |> fold_map (subst subs) ts
+ |-> (fn ts' => pair (list_comb (f, ts')))
+ | SOME (i', y) =>
+ fs
+ |> fold_map (subst subs) (xs @ ls @ rs')
+ ||> process_fun descr eqns (i', fname')
+ |-> (fn ts' => pair (list_comb (y, ts')))
+ end
+ else
+ fs
+ |> fold_map (subst subs) (f :: ts)
+ |-> (fn (f'::ts') => pair (list_comb (f', ts')))
+ end
+ | subst _ t fs = (t, fs);
+
+ (* translate rec equations into function arguments suitable for rec comb *)
+
+ fun trans eqns (cname, cargs) (fnames', fnss', fns) =
+ (case AList.lookup (op =) eqns cname of
+ NONE => (warning ("No equation for constructor " ^ quote cname ^
+ "\nin definition of function " ^ quote fname);
+ (fnames', fnss', (Const ("HOL.undefined", dummyT))::fns))
+ | SOME (ls, cargs', rs, rhs, eq) =>
+ let
+ val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
+ val rargs = map fst recs;
+ val subs = map (rpair dummyT o fst)
+ (rev (Term.rename_wrt_term rhs rargs));
+ val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
+ (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
+ handle PrimrecError (s, NONE) => primrec_error_eqn s eq
+ in (fnames'', fnss'',
+ (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns)
+ end)
+
+ in (case AList.lookup (op =) fnames i of
+ NONE =>
+ if exists (fn (_, v) => fname = v) fnames then
+ primrec_error ("inconsistent functions for datatype " ^ quote tname)
+ else
+ let
+ val (_, _, eqns) = the (AList.lookup (op =) eqns fname);
+ val (fnames', fnss', fns) = fold_rev (trans eqns) constrs
+ ((i, fname)::fnames, fnss, [])
+ in
+ (fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
+ end
+ | SOME fname' =>
+ if fname = fname' then (fnames, fnss)
+ else primrec_error ("inconsistent functions for datatype " ^ quote tname))
+ end;
+
+
+(* prepare functions needed for definitions *)
+
+fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
+ case AList.lookup (op =) fns i of
+ NONE =>
+ let
+ val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
+ replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
+ dummyT ---> HOLogic.unitT)) constrs;
+ val _ = warning ("No function definition for datatype " ^ quote tname)
+ in
+ (dummy_fns @ fs, defs)
+ end
+ | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs);
+
+
+(* make definition *)
+
+fun make_def ctxt fixes fs (fname, ls, rec_name, tname) =
+ let
+ val SOME (var, varT) = get_first (fn ((b, T), mx) =>
+ if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes;
+ val def_name = Thm.def_name (Long_Name.base_name fname);
+ val raw_rhs = fold_rev (fn T => fn t => Abs ("", T, t)) (map snd ls @ [dummyT])
+ (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1))))
+ val rhs = singleton (Syntax.check_terms ctxt)
+ (TypeInfer.constrain varT raw_rhs);
+ in (var, ((Binding.name def_name, []), rhs)) end;
+
+
+(* find datatypes which contain all datatypes in tnames' *)
+
+fun find_dts (dt_info : info Symtab.table) _ [] = []
+ | find_dts dt_info tnames' (tname::tnames) =
+ (case Symtab.lookup dt_info tname of
+ NONE => primrec_error (quote tname ^ " is not a datatype")
+ | SOME dt =>
+ if tnames' subset (map (#1 o snd) (#descr dt)) then
+ (tname, dt)::(find_dts dt_info tnames' tnames)
+ else find_dts dt_info tnames' tnames);
+
+
+(* distill primitive definition(s) from primrec specification *)
+
+fun distill lthy fixes eqs =
+ let
+ val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v
+ orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) eqs [];
+ val tnames = distinct (op =) (map (#1 o snd) eqns);
+ val dts = find_dts (Datatype.get_datatypes (ProofContext.theory_of lthy)) tnames tnames;
+ val main_fns = map (fn (tname, {index, ...}) =>
+ (index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts;
+ val {descr, rec_names, rec_rewrites, ...} =
+ if null dts then primrec_error
+ ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
+ else snd (hd dts);
+ val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []);
+ val (fs, raw_defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
+ val defs = map (make_def lthy fixes fs) raw_defs;
+ val names = map snd fnames;
+ val names_eqns = map fst eqns;
+ val _ = if gen_eq_set (op =) (names, names_eqns) then ()
+ else primrec_error ("functions " ^ commas_quote names_eqns ^
+ "\nare not mutually recursive");
+ val rec_rewrites' = map mk_meta_eq rec_rewrites;
+ val prefix = space_implode "_" (map (Long_Name.base_name o #1) raw_defs);
+ fun prove lthy defs =
+ let
+ val rewrites = rec_rewrites' @ map (snd o snd) defs;
+ fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1];
+ val _ = message ("Proving equations for primrec function(s) " ^ commas_quote names);
+ in map (fn eq => [Goal.prove lthy [] [] eq tac]) eqs end;
+ in ((prefix, (fs, defs)), prove) end
+ handle PrimrecError (msg, some_eqn) =>
+ error ("Primrec definition error:\n" ^ msg ^ (case some_eqn
+ of SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term lthy eqn)
+ | NONE => ""));
+
+
+(* primrec definition *)
+
+fun add_primrec_simple fixes ts lthy =
+ let
+ val ((prefix, (fs, defs)), prove) = distill lthy fixes ts;
+ in
+ lthy
+ |> fold_map (LocalTheory.define Thm.definitionK) defs
+ |-> (fn defs => `(fn lthy => (prefix, prove lthy defs)))
+ end;
+
+local
+
+fun gen_primrec set_group prep_spec raw_fixes raw_spec lthy =
+ let
+ val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
+ fun attr_bindings prefix = map (fn ((b, attrs), _) =>
+ (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec;
+ fun simp_attr_binding prefix = (Binding.qualify true prefix (Binding.name "simps"),
+ map (Attrib.internal o K)
+ [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add]);
+ in
+ lthy
+ |> set_group ? LocalTheory.set_group (serial_string ())
+ |> add_primrec_simple fixes (map snd spec)
+ |-> (fn (prefix, simps) => fold_map (LocalTheory.note Thm.generatedK)
+ (attr_bindings prefix ~~ simps)
+ #-> (fn simps' => LocalTheory.note Thm.generatedK
+ (simp_attr_binding prefix, maps snd simps')))
+ |>> snd
+ end;
+
+in
+
+val add_primrec = gen_primrec false Specification.check_spec;
+val add_primrec_cmd = gen_primrec true Specification.read_spec;
+
+end;
+
+fun add_primrec_global fixes specs thy =
+ let
+ val lthy = TheoryTarget.init NONE thy;
+ val (simps, lthy') = add_primrec fixes specs lthy;
+ val simps' = ProofContext.export lthy' lthy simps;
+ in (simps', LocalTheory.exit_global lthy') end;
+
+fun add_primrec_overloaded ops fixes specs thy =
+ let
+ val lthy = TheoryTarget.overloading ops thy;
+ val (simps, lthy') = add_primrec fixes specs lthy;
+ val simps' = ProofContext.export lthy' lthy simps;
+ in (simps', LocalTheory.exit_global lthy') end;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val opt_unchecked_name =
+ Scan.optional (P.$$$ "(" |-- P.!!!
+ (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
+ P.name >> pair false) --| P.$$$ ")")) (false, "");
+
+val old_primrec_decl =
+ opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop);
+
+val primrec_decl = P.opt_target -- P.fixes -- SpecParse.where_alt_specs;
+
+val _ =
+ OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
+ ((primrec_decl >> (fn ((opt_target, fixes), specs) =>
+ Toplevel.local_theory opt_target (add_primrec_cmd fixes specs #> snd)))
+ || (old_primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
+ Toplevel.theory (snd o
+ (if unchecked then OldPrimrec.add_primrec_unchecked else OldPrimrec.add_primrec)
+ alt_name (map P.triple_swap eqns)))));
+
+end;
+
+end;
--- a/src/HOL/Tools/primrec_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,332 +0,0 @@
-(* Title: HOL/Tools/primrec_package.ML
- Author: Stefan Berghofer, TU Muenchen; Norbert Voelker, FernUni Hagen;
- Florian Haftmann, TU Muenchen
-
-Package for defining functions on datatypes by primitive recursion.
-*)
-
-signature PRIMREC_PACKAGE =
-sig
- val add_primrec: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> local_theory -> thm list * local_theory
- val add_primrec_cmd: (binding * string option * mixfix) list ->
- (Attrib.binding * string) list -> local_theory -> thm list * local_theory
- val add_primrec_global: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> thm list * theory
- val add_primrec_overloaded: (string * (string * typ) * bool) list ->
- (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> thm list * theory
- val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
- local_theory -> (string * thm list list) * local_theory
-end;
-
-structure PrimrecPackage : PRIMREC_PACKAGE =
-struct
-
-open DatatypeAux;
-
-exception PrimrecError of string * term option;
-
-fun primrec_error msg = raise PrimrecError (msg, NONE);
-fun primrec_error_eqn msg eqn = raise PrimrecError (msg, SOME eqn);
-
-fun message s = if ! Toplevel.debug then tracing s else ();
-
-
-(* preprocessing of equations *)
-
-fun process_eqn is_fixed spec rec_fns =
- let
- val (vs, Ts) = split_list (strip_qnt_vars "all" spec);
- val body = strip_qnt_body "all" spec;
- val (vs', _) = Name.variants vs (Name.make_context (fold_aterms
- (fn Free (v, _) => insert (op =) v | _ => I) body []));
- val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body;
- val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn)
- handle TERM _ => primrec_error "not a proper equation";
- val (recfun, args) = strip_comb lhs;
- val fname = case recfun of Free (v, _) => if is_fixed v then v
- else primrec_error "illegal head of function equation"
- | _ => primrec_error "illegal head of function equation";
-
- val (ls', rest) = take_prefix is_Free args;
- val (middle, rs') = take_suffix is_Free rest;
- val rpos = length ls';
-
- val (constr, cargs') = if null middle then primrec_error "constructor missing"
- else strip_comb (hd middle);
- val (cname, T) = dest_Const constr
- handle TERM _ => primrec_error "ill-formed constructor";
- val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
- primrec_error "cannot determine datatype associated with function"
-
- val (ls, cargs, rs) =
- (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
- handle TERM _ => primrec_error "illegal argument in pattern";
- val lfrees = ls @ rs @ cargs;
-
- fun check_vars _ [] = ()
- | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn;
- in
- if length middle > 1 then
- primrec_error "more than one non-variable in pattern"
- else
- (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
- check_vars "extra variables on rhs: "
- (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
- |> filter_out (is_fixed o fst));
- case AList.lookup (op =) rec_fns fname of
- NONE =>
- (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))]))::rec_fns
- | SOME (_, rpos', eqns) =>
- if AList.defined (op =) eqns cname then
- primrec_error "constructor already occurred as pattern"
- else if rpos <> rpos' then
- primrec_error "position of recursive argument inconsistent"
- else
- AList.update (op =)
- (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn))::eqns))
- rec_fns)
- end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec;
-
-fun process_fun descr eqns (i, fname) (fnames, fnss) =
- let
- val (_, (tname, _, constrs)) = nth descr i;
-
- (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
-
- fun subst [] t fs = (t, fs)
- | subst subs (Abs (a, T, t)) fs =
- fs
- |> subst subs t
- |-> (fn t' => pair (Abs (a, T, t')))
- | subst subs (t as (_ $ _)) fs =
- let
- val (f, ts) = strip_comb t;
- in
- if is_Free f
- andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then
- let
- val (fname', _) = dest_Free f;
- val (_, rpos, _) = the (AList.lookup (op =) eqns fname');
- val (ls, rs) = chop rpos ts
- val (x', rs') = case rs
- of x' :: rs => (x', rs)
- | [] => primrec_error ("not enough arguments in recursive application\n"
- ^ "of function " ^ quote fname' ^ " on rhs");
- val (x, xs) = strip_comb x';
- in case AList.lookup (op =) subs x
- of NONE =>
- fs
- |> fold_map (subst subs) ts
- |-> (fn ts' => pair (list_comb (f, ts')))
- | SOME (i', y) =>
- fs
- |> fold_map (subst subs) (xs @ ls @ rs')
- ||> process_fun descr eqns (i', fname')
- |-> (fn ts' => pair (list_comb (y, ts')))
- end
- else
- fs
- |> fold_map (subst subs) (f :: ts)
- |-> (fn (f'::ts') => pair (list_comb (f', ts')))
- end
- | subst _ t fs = (t, fs);
-
- (* translate rec equations into function arguments suitable for rec comb *)
-
- fun trans eqns (cname, cargs) (fnames', fnss', fns) =
- (case AList.lookup (op =) eqns cname of
- NONE => (warning ("No equation for constructor " ^ quote cname ^
- "\nin definition of function " ^ quote fname);
- (fnames', fnss', (Const ("HOL.undefined", dummyT))::fns))
- | SOME (ls, cargs', rs, rhs, eq) =>
- let
- val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
- val rargs = map fst recs;
- val subs = map (rpair dummyT o fst)
- (rev (Term.rename_wrt_term rhs rargs));
- val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
- (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
- handle PrimrecError (s, NONE) => primrec_error_eqn s eq
- in (fnames'', fnss'',
- (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns)
- end)
-
- in (case AList.lookup (op =) fnames i of
- NONE =>
- if exists (fn (_, v) => fname = v) fnames then
- primrec_error ("inconsistent functions for datatype " ^ quote tname)
- else
- let
- val (_, _, eqns) = the (AList.lookup (op =) eqns fname);
- val (fnames', fnss', fns) = fold_rev (trans eqns) constrs
- ((i, fname)::fnames, fnss, [])
- in
- (fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
- end
- | SOME fname' =>
- if fname = fname' then (fnames, fnss)
- else primrec_error ("inconsistent functions for datatype " ^ quote tname))
- end;
-
-
-(* prepare functions needed for definitions *)
-
-fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
- case AList.lookup (op =) fns i of
- NONE =>
- let
- val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
- replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
- dummyT ---> HOLogic.unitT)) constrs;
- val _ = warning ("No function definition for datatype " ^ quote tname)
- in
- (dummy_fns @ fs, defs)
- end
- | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs);
-
-
-(* make definition *)
-
-fun make_def ctxt fixes fs (fname, ls, rec_name, tname) =
- let
- val SOME (var, varT) = get_first (fn ((b, T), mx) =>
- if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes;
- val def_name = Thm.def_name (Long_Name.base_name fname);
- val raw_rhs = fold_rev (fn T => fn t => Abs ("", T, t)) (map snd ls @ [dummyT])
- (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1))))
- val rhs = singleton (Syntax.check_terms ctxt)
- (TypeInfer.constrain varT raw_rhs);
- in (var, ((Binding.name def_name, []), rhs)) end;
-
-
-(* find datatypes which contain all datatypes in tnames' *)
-
-fun find_dts (dt_info : datatype_info Symtab.table) _ [] = []
- | find_dts dt_info tnames' (tname::tnames) =
- (case Symtab.lookup dt_info tname of
- NONE => primrec_error (quote tname ^ " is not a datatype")
- | SOME dt =>
- if tnames' subset (map (#1 o snd) (#descr dt)) then
- (tname, dt)::(find_dts dt_info tnames' tnames)
- else find_dts dt_info tnames' tnames);
-
-
-(* distill primitive definition(s) from primrec specification *)
-
-fun distill lthy fixes eqs =
- let
- val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v
- orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) eqs [];
- val tnames = distinct (op =) (map (#1 o snd) eqns);
- val dts = find_dts (DatatypePackage.get_datatypes (ProofContext.theory_of lthy)) tnames tnames;
- val main_fns = map (fn (tname, {index, ...}) =>
- (index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts;
- val {descr, rec_names, rec_rewrites, ...} =
- if null dts then primrec_error
- ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
- else snd (hd dts);
- val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []);
- val (fs, raw_defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
- val defs = map (make_def lthy fixes fs) raw_defs;
- val names = map snd fnames;
- val names_eqns = map fst eqns;
- val _ = if gen_eq_set (op =) (names, names_eqns) then ()
- else primrec_error ("functions " ^ commas_quote names_eqns ^
- "\nare not mutually recursive");
- val rec_rewrites' = map mk_meta_eq rec_rewrites;
- val prefix = space_implode "_" (map (Long_Name.base_name o #1) raw_defs);
- fun prove lthy defs =
- let
- val rewrites = rec_rewrites' @ map (snd o snd) defs;
- fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1];
- val _ = message ("Proving equations for primrec function(s) " ^ commas_quote names);
- in map (fn eq => [Goal.prove lthy [] [] eq tac]) eqs end;
- in ((prefix, (fs, defs)), prove) end
- handle PrimrecError (msg, some_eqn) =>
- error ("Primrec definition error:\n" ^ msg ^ (case some_eqn
- of SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term lthy eqn)
- | NONE => ""));
-
-
-(* primrec definition *)
-
-fun add_primrec_simple fixes ts lthy =
- let
- val ((prefix, (fs, defs)), prove) = distill lthy fixes ts;
- in
- lthy
- |> fold_map (LocalTheory.define Thm.definitionK) defs
- |-> (fn defs => `(fn lthy => (prefix, prove lthy defs)))
- end;
-
-local
-
-fun gen_primrec set_group prep_spec raw_fixes raw_spec lthy =
- let
- val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
- fun attr_bindings prefix = map (fn ((b, attrs), _) =>
- (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec;
- fun simp_attr_binding prefix = (Binding.qualify true prefix (Binding.name "simps"),
- map (Attrib.internal o K)
- [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add]);
- in
- lthy
- |> set_group ? LocalTheory.set_group (serial_string ())
- |> add_primrec_simple fixes (map snd spec)
- |-> (fn (prefix, simps) => fold_map (LocalTheory.note Thm.generatedK)
- (attr_bindings prefix ~~ simps)
- #-> (fn simps' => LocalTheory.note Thm.generatedK
- (simp_attr_binding prefix, maps snd simps')))
- |>> snd
- end;
-
-in
-
-val add_primrec = gen_primrec false Specification.check_spec;
-val add_primrec_cmd = gen_primrec true Specification.read_spec;
-
-end;
-
-fun add_primrec_global fixes specs thy =
- let
- val lthy = TheoryTarget.init NONE thy;
- val (simps, lthy') = add_primrec fixes specs lthy;
- val simps' = ProofContext.export lthy' lthy simps;
- in (simps', LocalTheory.exit_global lthy') end;
-
-fun add_primrec_overloaded ops fixes specs thy =
- let
- val lthy = TheoryTarget.overloading ops thy;
- val (simps, lthy') = add_primrec fixes specs lthy;
- val simps' = ProofContext.export lthy' lthy simps;
- in (simps', LocalTheory.exit_global lthy') end;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_unchecked_name =
- Scan.optional (P.$$$ "(" |-- P.!!!
- (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
- P.name >> pair false) --| P.$$$ ")")) (false, "");
-
-val old_primrec_decl =
- opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop);
-
-val primrec_decl = P.opt_target -- P.fixes -- SpecParse.where_alt_specs;
-
-val _ =
- OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
- ((primrec_decl >> (fn ((opt_target, fixes), specs) =>
- Toplevel.local_theory opt_target (add_primrec_cmd fixes specs #> snd)))
- || (old_primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
- Toplevel.theory (snd o
- (if unchecked then OldPrimrecPackage.add_primrec_unchecked else OldPrimrecPackage.add_primrec)
- alt_name (map P.triple_swap eqns)))));
-
-end;
-
-end;
--- a/src/HOL/Tools/quickcheck_generators.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/quickcheck_generators.ML Tue Jun 23 15:48:55 2009 +0100
@@ -12,10 +12,10 @@
-> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
val ensure_random_typecopy: string -> theory -> theory
val random_aux_specification: string -> term list -> local_theory -> local_theory
- val mk_random_aux_eqs: theory -> DatatypeAux.descr -> (string * sort) list
+ val mk_random_aux_eqs: theory -> Datatype.descr -> (string * sort) list
-> string list -> string list * string list -> typ list * typ list
-> string * (term list * (term * term) list)
- val ensure_random_datatype: DatatypeAux.datatype_config -> string list -> theory -> theory
+ val ensure_random_datatype: Datatype.config -> string list -> theory -> theory
val eval_ref: (unit -> int -> seed -> term list option * seed) option ref
val setup: theory -> theory
end;
@@ -122,7 +122,7 @@
fun ensure_random_typecopy tyco thy =
let
val SOME { vs = raw_vs, constr, typ = raw_typ, ... } =
- TypecopyPackage.get_info thy tyco;
+ Typecopy.get_info thy tyco;
val constrain = curry (Sorts.inter_sort (Sign.classes_of thy));
val typ = map_atyps (fn TFree (v, sort) =>
TFree (v, constrain sort @{sort random})) raw_typ;
@@ -168,7 +168,7 @@
val t_rhs = lambda t_k proto_t_rhs;
val eqs0 = [subst_v @{term "0::code_numeral"} eq, subst_v (@{term "Suc_code_numeral"} $ t_k) eq];
val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0;
- val ((_, eqs2), lthy') = PrimrecPackage.add_primrec_simple
+ val ((_, eqs2), lthy') = Primrec.add_primrec_simple
[((Binding.name random_aux, T), NoSyn)] eqs1 lthy;
val eq_tac = ALLGOALS (simp_tac rew_ss)
THEN (ALLGOALS (ProofContext.fact_tac (flat eqs2)));
@@ -361,17 +361,21 @@
val pp = Syntax.pp_global thy;
val algebra = Sign.classes_of thy;
val (descr, raw_vs, tycos, (names, auxnames), raw_TUs) =
- DatatypePackage.the_datatype_descr thy raw_tycos;
+ Datatype.the_datatype_descr thy raw_tycos;
+ val typrep_vs = (map o apsnd)
+ (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs;
val random_insts = (map (rpair @{sort random}) o flat o maps snd o maps snd)
- (DatatypeAux.interpret_construction descr raw_vs { atyp = single, dtyp = (K o K o K) [] });
+ (DatatypeAux.interpret_construction descr typrep_vs
+ { atyp = single, dtyp = (K o K o K) [] });
val term_of_insts = (map (rpair @{sort term_of}) o flat o maps snd o maps snd)
- (DatatypeAux.interpret_construction descr raw_vs { atyp = K [], dtyp = K o K });
+ (DatatypeAux.interpret_construction descr typrep_vs
+ { atyp = K [], dtyp = K o K });
val has_inst = exists (fn tyco =>
can (Sorts.mg_domain algebra tyco) @{sort random}) tycos;
in if has_inst then thy
- else case perhaps_constrain thy (random_insts @ term_of_insts) raw_vs
+ else case perhaps_constrain thy (random_insts @ term_of_insts) typrep_vs
of SOME constrain => mk_random_datatype config descr
- (map constrain raw_vs) tycos (names, auxnames)
+ (map constrain typrep_vs) tycos (names, auxnames)
((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy
| NONE => thy
end;
@@ -381,7 +385,7 @@
val setup = Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
#> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of)
- #> TypecopyPackage.interpretation ensure_random_typecopy
- #> DatatypePackage.interpretation ensure_random_datatype;
+ #> Typecopy.interpretation ensure_random_typecopy
+ #> Datatype.interpretation ensure_random_datatype;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/recdef.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,331 @@
+(* Title: HOL/Tools/recdef.ML
+ Author: Markus Wenzel, TU Muenchen
+
+Wrapper module for Konrad Slind's TFL package.
+*)
+
+signature RECDEF =
+sig
+ val get_recdef: theory -> string
+ -> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
+ val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
+ val simp_add: attribute
+ val simp_del: attribute
+ val cong_add: attribute
+ val cong_del: attribute
+ val wf_add: attribute
+ val wf_del: attribute
+ val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
+ Attrib.src option -> theory -> theory
+ * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
+ val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
+ theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
+ val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
+ -> theory -> theory * {induct_rules: thm}
+ val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
+ val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool ->
+ local_theory -> Proof.state
+ val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool ->
+ local_theory -> Proof.state
+ val setup: theory -> theory
+end;
+
+structure Recdef: RECDEF =
+struct
+
+
+(** recdef hints **)
+
+(* type hints *)
+
+type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
+
+fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
+fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
+
+fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
+fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
+fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
+
+fun pretty_hints ({simps, congs, wfs}: hints) =
+ [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps),
+ Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)),
+ Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)];
+
+
+(* congruence rules *)
+
+local
+
+val cong_head =
+ fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
+
+fun prep_cong raw_thm =
+ let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
+
+in
+
+fun add_cong raw_thm congs =
+ let
+ val (c, thm) = prep_cong raw_thm;
+ val _ = if AList.defined (op =) congs c
+ then warning ("Overwriting recdef congruence rule for " ^ quote c)
+ else ();
+ in AList.update (op =) (c, thm) congs end;
+
+fun del_cong raw_thm congs =
+ let
+ val (c, thm) = prep_cong raw_thm;
+ val _ = if AList.defined (op =) congs c
+ then ()
+ else warning ("No recdef congruence rule for " ^ quote c);
+ in AList.delete (op =) c congs end;
+
+end;
+
+
+
+(** global and local recdef data **)
+
+(* theory data *)
+
+type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
+
+structure GlobalRecdefData = TheoryDataFun
+(
+ type T = recdef_info Symtab.table * hints;
+ val empty = (Symtab.empty, mk_hints ([], [], [])): T;
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
+ (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
+ (Symtab.merge (K true) (tab1, tab2),
+ mk_hints (Thm.merge_thms (simps1, simps2),
+ AList.merge (op =) Thm.eq_thm (congs1, congs2),
+ Thm.merge_thms (wfs1, wfs2)));
+);
+
+val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
+
+fun put_recdef name info thy =
+ let
+ val (tab, hints) = GlobalRecdefData.get thy;
+ val tab' = Symtab.update_new (name, info) tab
+ handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
+ in GlobalRecdefData.put (tab', hints) thy end;
+
+val get_global_hints = #2 o GlobalRecdefData.get;
+
+
+(* proof data *)
+
+structure LocalRecdefData = ProofDataFun
+(
+ type T = hints;
+ val init = get_global_hints;
+);
+
+val get_hints = LocalRecdefData.get;
+fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
+
+
+(* attributes *)
+
+fun attrib f = Thm.declaration_attribute (map_hints o f);
+
+val simp_add = attrib (map_simps o Thm.add_thm);
+val simp_del = attrib (map_simps o Thm.del_thm);
+val cong_add = attrib (map_congs o add_cong);
+val cong_del = attrib (map_congs o del_cong);
+val wf_add = attrib (map_wfs o Thm.add_thm);
+val wf_del = attrib (map_wfs o Thm.del_thm);
+
+
+(* modifiers *)
+
+val recdef_simpN = "recdef_simp";
+val recdef_congN = "recdef_cong";
+val recdef_wfN = "recdef_wf";
+
+val recdef_modifiers =
+ [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
+ Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
+ Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
+ Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
+ Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
+ Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
+ Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
+ Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
+ Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
+ Clasimp.clasimp_modifiers;
+
+
+
+(** prepare_hints(_i) **)
+
+fun prepare_hints thy opt_src =
+ let
+ val ctxt0 = ProofContext.init thy;
+ val ctxt =
+ (case opt_src of
+ NONE => ctxt0
+ | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0));
+ val {simps, congs, wfs} = get_hints ctxt;
+ val cs = local_claset_of ctxt;
+ val ss = local_simpset_of ctxt addsimps simps;
+ in (cs, ss, rev (map snd congs), wfs) end;
+
+fun prepare_hints_i thy () =
+ let
+ val ctxt0 = ProofContext.init thy;
+ val {simps, congs, wfs} = get_global_hints thy;
+ in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;
+
+
+
+(** add_recdef(_i) **)
+
+fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
+
+fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
+ let
+ val _ = requires_recdef thy;
+
+ val name = Sign.intern_const thy raw_name;
+ val bname = Long_Name.base_name name;
+ val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
+
+ val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
+ val eq_atts = map (map (prep_att thy)) raw_eq_atts;
+
+ val (cs, ss, congs, wfs) = prep_hints thy hints;
+ (*We must remove imp_cong to prevent looping when the induction rule
+ is simplified. Many induction rules have nested implications that would
+ give rise to looping conditional rewriting.*)
+ val (thy, {rules = rules_idx, induct, tcs}) =
+ tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
+ congs wfs name R eqs;
+ val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
+ val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
+ Code.add_default_eqn_attribute, Quickcheck_RecFun_Simp_Thms.add] else [];
+
+ val ((simps' :: rules', [induct']), thy) =
+ thy
+ |> Sign.add_path bname
+ |> PureThy.add_thmss
+ (((Binding.name "simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
+ ||>> PureThy.add_thms [((Binding.name "induct", induct), [])];
+ val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
+ val thy =
+ thy
+ |> put_recdef name result
+ |> Sign.parent_path;
+ in (thy, result) end;
+
+val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
+fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
+
+
+
+(** defer_recdef(_i) **)
+
+fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =
+ let
+ val name = Sign.intern_const thy raw_name;
+ val bname = Long_Name.base_name name;
+
+ val _ = requires_recdef thy;
+ val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
+
+ val congs = eval_thms (ProofContext.init thy) raw_congs;
+ val (thy2, induct_rules) = tfl_fn thy congs name eqs;
+ val ([induct_rules'], thy3) =
+ thy2
+ |> Sign.add_path bname
+ |> PureThy.add_thms [((Binding.name "induct_rules", induct_rules), [])]
+ ||> Sign.parent_path;
+ in (thy3, {induct_rules = induct_rules'}) end;
+
+val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms;
+val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);
+
+
+
+(** recdef_tc(_i) **)
+
+fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val name = prep_name thy raw_name;
+ val atts = map (prep_att thy) raw_atts;
+ val tcs =
+ (case get_recdef thy name of
+ NONE => error ("No recdef definition of constant: " ^ quote name)
+ | SOME {tcs, ...} => tcs);
+ val i = the_default 1 opt_i;
+ val tc = nth tcs (i - 1) handle Subscript =>
+ error ("No termination condition #" ^ string_of_int i ^
+ " in recdef definition of " ^ quote name);
+ in
+ Specification.theorem Thm.internalK NONE (K I) (Binding.name bname, atts)
+ [] (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
+ end;
+
+val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
+val recdef_tc_i = gen_recdef_tc (K I) (K I);
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
+ "declaration of recdef simp rule" #>
+ Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
+ "declaration of recdef cong rule" #>
+ Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
+ "declaration of recdef wf rule";
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = List.app OuterKeyword.keyword ["permissive", "congs", "hints"];
+
+val hints =
+ P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- Args.parse) --| P.$$$ ")") >> Args.src;
+
+val recdef_decl =
+ Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
+ P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
+ -- Scan.option hints
+ >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
+
+val _ =
+ OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
+ (recdef_decl >> Toplevel.theory);
+
+
+val defer_recdef_decl =
+ P.name -- Scan.repeat1 P.prop --
+ Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (SpecParse.xthms1 --| P.$$$ ")")) []
+ >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
+
+val _ =
+ OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
+ (defer_recdef_decl >> Toplevel.theory);
+
+val _ =
+ OuterSyntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
+ K.thy_goal
+ ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.xname --
+ Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
+ >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
+
+end;
+
+end;
--- a/src/HOL/Tools/recdef_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,331 +0,0 @@
-(* Title: HOL/Tools/recdef_package.ML
- Author: Markus Wenzel, TU Muenchen
-
-Wrapper module for Konrad Slind's TFL package.
-*)
-
-signature RECDEF_PACKAGE =
-sig
- val get_recdef: theory -> string
- -> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
- val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
- val simp_add: attribute
- val simp_del: attribute
- val cong_add: attribute
- val cong_del: attribute
- val wf_add: attribute
- val wf_del: attribute
- val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
- Attrib.src option -> theory -> theory
- * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
- val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
- theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
- val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
- -> theory -> theory * {induct_rules: thm}
- val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
- val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool ->
- local_theory -> Proof.state
- val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool ->
- local_theory -> Proof.state
- val setup: theory -> theory
-end;
-
-structure RecdefPackage: RECDEF_PACKAGE =
-struct
-
-
-(** recdef hints **)
-
-(* type hints *)
-
-type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
-
-fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
-fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
-
-fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
-fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
-fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
-
-fun pretty_hints ({simps, congs, wfs}: hints) =
- [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps),
- Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)),
- Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)];
-
-
-(* congruence rules *)
-
-local
-
-val cong_head =
- fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
-
-fun prep_cong raw_thm =
- let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
-
-in
-
-fun add_cong raw_thm congs =
- let
- val (c, thm) = prep_cong raw_thm;
- val _ = if AList.defined (op =) congs c
- then warning ("Overwriting recdef congruence rule for " ^ quote c)
- else ();
- in AList.update (op =) (c, thm) congs end;
-
-fun del_cong raw_thm congs =
- let
- val (c, thm) = prep_cong raw_thm;
- val _ = if AList.defined (op =) congs c
- then ()
- else warning ("No recdef congruence rule for " ^ quote c);
- in AList.delete (op =) c congs end;
-
-end;
-
-
-
-(** global and local recdef data **)
-
-(* theory data *)
-
-type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
-
-structure GlobalRecdefData = TheoryDataFun
-(
- type T = recdef_info Symtab.table * hints;
- val empty = (Symtab.empty, mk_hints ([], [], [])): T;
- val copy = I;
- val extend = I;
- fun merge _
- ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
- (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
- (Symtab.merge (K true) (tab1, tab2),
- mk_hints (Thm.merge_thms (simps1, simps2),
- AList.merge (op =) Thm.eq_thm (congs1, congs2),
- Thm.merge_thms (wfs1, wfs2)));
-);
-
-val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
-
-fun put_recdef name info thy =
- let
- val (tab, hints) = GlobalRecdefData.get thy;
- val tab' = Symtab.update_new (name, info) tab
- handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
- in GlobalRecdefData.put (tab', hints) thy end;
-
-val get_global_hints = #2 o GlobalRecdefData.get;
-
-
-(* proof data *)
-
-structure LocalRecdefData = ProofDataFun
-(
- type T = hints;
- val init = get_global_hints;
-);
-
-val get_hints = LocalRecdefData.get;
-fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
-
-
-(* attributes *)
-
-fun attrib f = Thm.declaration_attribute (map_hints o f);
-
-val simp_add = attrib (map_simps o Thm.add_thm);
-val simp_del = attrib (map_simps o Thm.del_thm);
-val cong_add = attrib (map_congs o add_cong);
-val cong_del = attrib (map_congs o del_cong);
-val wf_add = attrib (map_wfs o Thm.add_thm);
-val wf_del = attrib (map_wfs o Thm.del_thm);
-
-
-(* modifiers *)
-
-val recdef_simpN = "recdef_simp";
-val recdef_congN = "recdef_cong";
-val recdef_wfN = "recdef_wf";
-
-val recdef_modifiers =
- [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
- Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
- Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
- Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
- Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
- Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
- Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
- Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
- Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
- Clasimp.clasimp_modifiers;
-
-
-
-(** prepare_hints(_i) **)
-
-fun prepare_hints thy opt_src =
- let
- val ctxt0 = ProofContext.init thy;
- val ctxt =
- (case opt_src of
- NONE => ctxt0
- | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0));
- val {simps, congs, wfs} = get_hints ctxt;
- val cs = local_claset_of ctxt;
- val ss = local_simpset_of ctxt addsimps simps;
- in (cs, ss, rev (map snd congs), wfs) end;
-
-fun prepare_hints_i thy () =
- let
- val ctxt0 = ProofContext.init thy;
- val {simps, congs, wfs} = get_global_hints thy;
- in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;
-
-
-
-(** add_recdef(_i) **)
-
-fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
-
-fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
- let
- val _ = requires_recdef thy;
-
- val name = Sign.intern_const thy raw_name;
- val bname = Long_Name.base_name name;
- val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
-
- val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
- val eq_atts = map (map (prep_att thy)) raw_eq_atts;
-
- val (cs, ss, congs, wfs) = prep_hints thy hints;
- (*We must remove imp_cong to prevent looping when the induction rule
- is simplified. Many induction rules have nested implications that would
- give rise to looping conditional rewriting.*)
- val (thy, {rules = rules_idx, induct, tcs}) =
- tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
- congs wfs name R eqs;
- val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
- val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
- Code.add_default_eqn_attribute, Quickcheck_RecFun_Simp_Thms.add] else [];
-
- val ((simps' :: rules', [induct']), thy) =
- thy
- |> Sign.add_path bname
- |> PureThy.add_thmss
- (((Binding.name "simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
- ||>> PureThy.add_thms [((Binding.name "induct", induct), [])];
- val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
- val thy =
- thy
- |> put_recdef name result
- |> Sign.parent_path;
- in (thy, result) end;
-
-val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
-fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
-
-
-
-(** defer_recdef(_i) **)
-
-fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =
- let
- val name = Sign.intern_const thy raw_name;
- val bname = Long_Name.base_name name;
-
- val _ = requires_recdef thy;
- val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
-
- val congs = eval_thms (ProofContext.init thy) raw_congs;
- val (thy2, induct_rules) = tfl_fn thy congs name eqs;
- val ([induct_rules'], thy3) =
- thy2
- |> Sign.add_path bname
- |> PureThy.add_thms [((Binding.name "induct_rules", induct_rules), [])]
- ||> Sign.parent_path;
- in (thy3, {induct_rules = induct_rules'}) end;
-
-val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms;
-val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);
-
-
-
-(** recdef_tc(_i) **)
-
-fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =
- let
- val thy = ProofContext.theory_of lthy;
- val name = prep_name thy raw_name;
- val atts = map (prep_att thy) raw_atts;
- val tcs =
- (case get_recdef thy name of
- NONE => error ("No recdef definition of constant: " ^ quote name)
- | SOME {tcs, ...} => tcs);
- val i = the_default 1 opt_i;
- val tc = nth tcs (i - 1) handle Subscript =>
- error ("No termination condition #" ^ string_of_int i ^
- " in recdef definition of " ^ quote name);
- in
- Specification.theorem Thm.internalK NONE (K I) (Binding.name bname, atts)
- [] (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
- end;
-
-val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
-val recdef_tc_i = gen_recdef_tc (K I) (K I);
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
- "declaration of recdef simp rule" #>
- Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
- "declaration of recdef cong rule" #>
- Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
- "declaration of recdef wf rule";
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword ["permissive", "congs", "hints"];
-
-val hints =
- P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- Args.parse) --| P.$$$ ")") >> Args.src;
-
-val recdef_decl =
- Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
- P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
- -- Scan.option hints
- >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
-
-val _ =
- OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
- (recdef_decl >> Toplevel.theory);
-
-
-val defer_recdef_decl =
- P.name -- Scan.repeat1 P.prop --
- Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (SpecParse.xthms1 --| P.$$$ ")")) []
- >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
-
-val _ =
- OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
- (defer_recdef_decl >> Toplevel.theory);
-
-val _ =
- OuterSyntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
- K.thy_goal
- ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.xname --
- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
- >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
-
-end;
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/record.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,2325 @@
+(* Title: HOL/Tools/record.ML
+ Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
+
+Extensible records with structural subtyping in HOL.
+*)
+
+
+signature BASIC_RECORD =
+sig
+ val record_simproc: simproc
+ val record_eq_simproc: simproc
+ val record_upd_simproc: simproc
+ val record_split_simproc: (term -> int) -> simproc
+ val record_ex_sel_eq_simproc: simproc
+ val record_split_tac: int -> tactic
+ val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic
+ val record_split_name: string
+ val record_split_wrapper: string * wrapper
+ val print_record_type_abbr: bool ref
+ val print_record_type_as_fields: bool ref
+end;
+
+signature RECORD =
+sig
+ include BASIC_RECORD
+ val timing: bool ref
+ val record_quick_and_dirty_sensitive: bool ref
+ val updateN: string
+ val updN: string
+ val ext_typeN: string
+ val extN: string
+ val makeN: string
+ val moreN: string
+ val ext_dest: string
+
+ val last_extT: typ -> (string * typ list) option
+ val dest_recTs : typ -> (string * typ list) list
+ val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ)
+ val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ)
+ val get_parent: theory -> string -> (typ list * string) option
+ val get_extension: theory -> string -> (string * typ list) option
+ val get_extinjects: theory -> thm list
+ val get_simpset: theory -> simpset
+ val print_records: theory -> unit
+ val read_typ: Proof.context -> string -> (string * sort) list -> typ * (string * sort) list
+ val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list
+ val add_record: bool -> string list * string -> string option -> (string * string * mixfix) list
+ -> theory -> theory
+ val add_record_i: bool -> string list * string -> (typ list * string) option
+ -> (string * typ * mixfix) list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+
+structure Record: RECORD =
+struct
+
+val eq_reflection = thm "eq_reflection";
+val rec_UNIV_I = thm "rec_UNIV_I";
+val rec_True_simp = thm "rec_True_simp";
+val Pair_eq = thm "Product_Type.prod.inject";
+val atomize_all = thm "HOL.atomize_all";
+val atomize_imp = thm "HOL.atomize_imp";
+val meta_allE = thm "Pure.meta_allE";
+val prop_subst = thm "prop_subst";
+val Pair_sel_convs = [fst_conv,snd_conv];
+val K_record_comp = @{thm "K_record_comp"};
+val K_comp_convs = [@{thm o_apply}, K_record_comp]
+
+(** name components **)
+
+val rN = "r";
+val wN = "w";
+val moreN = "more";
+val schemeN = "_scheme";
+val ext_typeN = "_ext_type";
+val extN ="_ext";
+val casesN = "_cases";
+val ext_dest = "_sel";
+val updateN = "_update";
+val updN = "_upd";
+val makeN = "make";
+val fields_selN = "fields";
+val extendN = "extend";
+val truncateN = "truncate";
+
+(*see typedef.ML*)
+val RepN = "Rep_";
+val AbsN = "Abs_";
+
+(*** utilities ***)
+
+fun but_last xs = fst (split_last xs);
+
+fun varifyT midx =
+ let fun varify (a, S) = TVar ((a, midx + 1), S);
+ in map_type_tfree varify end;
+
+fun domain_type' T =
+ domain_type T handle Match => T;
+
+fun range_type' T =
+ range_type T handle Match => T;
+
+(* messages *)
+
+fun trace_thm str thm =
+ tracing (str ^ (Pretty.string_of (Display.pretty_thm thm)));
+
+fun trace_thms str thms =
+ (tracing str; map (trace_thm "") thms);
+
+fun trace_term str t =
+ tracing (str ^ Syntax.string_of_term_global Pure.thy t);
+
+(* timing *)
+
+val timing = ref false;
+fun timeit_msg s x = if !timing then (warning s; timeit x) else x ();
+fun timing_msg s = if !timing then warning s else ();
+
+(* syntax *)
+
+fun prune n xs = Library.drop (n, xs);
+fun prefix_base s = Long_Name.map_base_name (fn bname => s ^ bname);
+
+val Trueprop = HOLogic.mk_Trueprop;
+fun All xs t = Term.list_all_free (xs, t);
+
+infix 9 $$;
+infix 0 :== ===;
+infixr 0 ==>;
+
+val (op $$) = Term.list_comb;
+val (op :==) = PrimitiveDefs.mk_defpair;
+val (op ===) = Trueprop o HOLogic.mk_eq;
+val (op ==>) = Logic.mk_implies;
+
+(* morphisms *)
+
+fun mk_RepN name = suffix ext_typeN (prefix_base RepN name);
+fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name);
+
+fun mk_Rep name repT absT =
+ Const (suffix ext_typeN (prefix_base RepN name),absT --> repT);
+
+fun mk_Abs name repT absT =
+ Const (mk_AbsN name,repT --> absT);
+
+(* constructor *)
+
+fun mk_extC (name,T) Ts = (suffix extN name, Ts ---> T);
+
+fun mk_ext (name,T) ts =
+ let val Ts = map fastype_of ts
+ in list_comb (Const (mk_extC (name,T) Ts),ts) end;
+
+(* cases *)
+
+fun mk_casesC (name,T,vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT)
+
+fun mk_cases (name,T,vT) f =
+ let val Ts = binder_types (fastype_of f)
+ in Const (mk_casesC (name,T,vT) Ts) $ f end;
+
+(* selector *)
+
+fun mk_selC sT (c,T) = (c,sT --> T);
+
+fun mk_sel s (c,T) =
+ let val sT = fastype_of s
+ in Const (mk_selC sT (c,T)) $ s end;
+
+(* updates *)
+
+fun mk_updC sfx sT (c,T) = (suffix sfx c, (T --> T) --> sT --> sT);
+
+fun mk_upd' sfx c v sT =
+ let val vT = domain_type (fastype_of v);
+ in Const (mk_updC sfx sT (c, vT)) $ v end;
+
+fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s
+
+(* types *)
+
+fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) =
+ (case try (unsuffix ext_typeN) c_ext_type of
+ NONE => raise TYPE ("Record.dest_recT", [typ], [])
+ | SOME c => ((c, Ts), List.last Ts))
+ | dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []);
+
+fun is_recT T =
+ (case try dest_recT T of NONE => false | SOME _ => true);
+
+fun dest_recTs T =
+ let val ((c, Ts), U) = dest_recT T
+ in (c, Ts) :: dest_recTs U
+ end handle TYPE _ => [];
+
+fun last_extT T =
+ let val ((c, Ts), U) = dest_recT T
+ in (case last_extT U of
+ NONE => SOME (c,Ts)
+ | SOME l => SOME l)
+ end handle TYPE _ => NONE
+
+fun rec_id i T =
+ let val rTs = dest_recTs T
+ val rTs' = if i < 0 then rTs else Library.take (i,rTs)
+ in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end;
+
+(*** extend theory by record definition ***)
+
+(** record info **)
+
+(* type record_info and parent_info *)
+
+type record_info =
+ {args: (string * sort) list,
+ parent: (typ list * string) option,
+ fields: (string * typ) list,
+ extension: (string * typ list),
+ induct: thm
+ };
+
+fun make_record_info args parent fields extension induct =
+ {args = args, parent = parent, fields = fields, extension = extension,
+ induct = induct}: record_info;
+
+
+type parent_info =
+ {name: string,
+ fields: (string * typ) list,
+ extension: (string * typ list),
+ induct: thm
+};
+
+fun make_parent_info name fields extension induct =
+ {name = name, fields = fields, extension = extension, induct = induct}: parent_info;
+
+
+(* theory data *)
+
+type record_data =
+ {records: record_info Symtab.table,
+ sel_upd:
+ {selectors: unit Symtab.table,
+ updates: string Symtab.table,
+ simpset: Simplifier.simpset},
+ equalities: thm Symtab.table,
+ extinjects: thm list,
+ extsplit: thm Symtab.table, (* maps extension name to split rule *)
+ splits: (thm*thm*thm*thm) Symtab.table, (* !!,!,EX - split-equalities,induct rule *)
+ extfields: (string*typ) list Symtab.table, (* maps extension to its fields *)
+ fieldext: (string*typ list) Symtab.table (* maps field to its extension *)
+};
+
+fun make_record_data
+ records sel_upd equalities extinjects extsplit splits extfields fieldext =
+ {records = records, sel_upd = sel_upd,
+ equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
+ extfields = extfields, fieldext = fieldext }: record_data;
+
+structure RecordsData = TheoryDataFun
+(
+ type T = record_data;
+ val empty =
+ make_record_data Symtab.empty
+ {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss}
+ Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty;
+
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ({records = recs1,
+ sel_upd = {selectors = sels1, updates = upds1, simpset = ss1},
+ equalities = equalities1,
+ extinjects = extinjects1,
+ extsplit = extsplit1,
+ splits = splits1,
+ extfields = extfields1,
+ fieldext = fieldext1},
+ {records = recs2,
+ sel_upd = {selectors = sels2, updates = upds2, simpset = ss2},
+ equalities = equalities2,
+ extinjects = extinjects2,
+ extsplit = extsplit2,
+ splits = splits2,
+ extfields = extfields2,
+ fieldext = fieldext2}) =
+ make_record_data
+ (Symtab.merge (K true) (recs1, recs2))
+ {selectors = Symtab.merge (K true) (sels1, sels2),
+ updates = Symtab.merge (K true) (upds1, upds2),
+ simpset = Simplifier.merge_ss (ss1, ss2)}
+ (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2))
+ (Library.merge Thm.eq_thm_prop (extinjects1, extinjects2))
+ (Symtab.merge Thm.eq_thm_prop (extsplit1,extsplit2))
+ (Symtab.merge (fn ((a,b,c,d),(w,x,y,z))
+ => Thm.eq_thm (a,w) andalso Thm.eq_thm (b,x) andalso
+ Thm.eq_thm (c,y) andalso Thm.eq_thm (d,z))
+ (splits1, splits2))
+ (Symtab.merge (K true) (extfields1,extfields2))
+ (Symtab.merge (K true) (fieldext1,fieldext2));
+);
+
+fun print_records thy =
+ let
+ val {records = recs, ...} = RecordsData.get thy;
+ val prt_typ = Syntax.pretty_typ_global thy;
+
+ fun pretty_parent NONE = []
+ | pretty_parent (SOME (Ts, name)) =
+ [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
+
+ fun pretty_field (c, T) = Pretty.block
+ [Pretty.str (Sign.extern_const thy c), Pretty.str " ::",
+ Pretty.brk 1, Pretty.quote (prt_typ T)];
+
+ fun pretty_record (name, {args, parent, fields, ...}: record_info) =
+ Pretty.block (Pretty.fbreaks (Pretty.block
+ [prt_typ (Type (name, map TFree args)), Pretty.str " = "] ::
+ pretty_parent parent @ map pretty_field fields));
+ in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end;
+
+
+(* access 'records' *)
+
+val get_record = Symtab.lookup o #records o RecordsData.get;
+
+fun put_record name info thy =
+ let
+ val {records, sel_upd, equalities, extinjects,extsplit,splits,extfields,fieldext} =
+ RecordsData.get thy;
+ val data = make_record_data (Symtab.update (name, info) records)
+ sel_upd equalities extinjects extsplit splits extfields fieldext;
+ in RecordsData.put data thy end;
+
+
+(* access 'sel_upd' *)
+
+val get_sel_upd = #sel_upd o RecordsData.get;
+
+val is_selector = Symtab.defined o #selectors o get_sel_upd;
+val get_updates = Symtab.lookup o #updates o get_sel_upd;
+fun get_simpset thy = Simplifier.theory_context thy (#simpset (get_sel_upd thy));
+
+fun put_sel_upd names simps = RecordsData.map (fn {records,
+ sel_upd = {selectors, updates, simpset},
+ equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_record_data records
+ {selectors = fold (fn name => Symtab.update (name, ())) names selectors,
+ updates = fold (fn name => Symtab.update ((suffix updateN) name, name)) names updates,
+ simpset = Simplifier.addsimps (simpset, simps)}
+ equalities extinjects extsplit splits extfields fieldext);
+
+
+(* access 'equalities' *)
+
+fun add_record_equalities name thm thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
+ RecordsData.get thy;
+ val data = make_record_data records sel_upd
+ (Symtab.update_new (name, thm) equalities) extinjects extsplit
+ splits extfields fieldext;
+ in RecordsData.put data thy end;
+
+val get_equalities =Symtab.lookup o #equalities o RecordsData.get;
+
+
+(* access 'extinjects' *)
+
+fun add_extinjects thm thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
+ RecordsData.get thy;
+ val data =
+ make_record_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects) extsplit
+ splits extfields fieldext;
+ in RecordsData.put data thy end;
+
+val get_extinjects = rev o #extinjects o RecordsData.get;
+
+
+(* access 'extsplit' *)
+
+fun add_extsplit name thm thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
+ RecordsData.get thy;
+ val data = make_record_data records sel_upd
+ equalities extinjects (Symtab.update_new (name, thm) extsplit) splits
+ extfields fieldext;
+ in RecordsData.put data thy end;
+
+val get_extsplit = Symtab.lookup o #extsplit o RecordsData.get;
+
+
+(* access 'splits' *)
+
+fun add_record_splits name thmP thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
+ RecordsData.get thy;
+ val data = make_record_data records sel_upd
+ equalities extinjects extsplit (Symtab.update_new (name, thmP) splits)
+ extfields fieldext;
+ in RecordsData.put data thy end;
+
+val get_splits = Symtab.lookup o #splits o RecordsData.get;
+
+
+(* parent/extension of named record *)
+
+val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o RecordsData.get);
+val get_extension = Option.map #extension oo (Symtab.lookup o #records o RecordsData.get);
+
+
+(* access 'extfields' *)
+
+fun add_extfields name fields thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit,splits, extfields, fieldext} =
+ RecordsData.get thy;
+ val data = make_record_data records sel_upd
+ equalities extinjects extsplit splits
+ (Symtab.update_new (name, fields) extfields) fieldext;
+ in RecordsData.put data thy end;
+
+val get_extfields = Symtab.lookup o #extfields o RecordsData.get;
+
+fun get_extT_fields thy T =
+ let
+ val ((name,Ts),moreT) = dest_recT T;
+ val recname = let val (nm::recn::rst) = rev (Long_Name.explode name)
+ in Long_Name.implode (rev (nm::rst)) end;
+ val midx = maxidx_of_typs (moreT::Ts);
+ val varifyT = varifyT midx;
+ val {records,extfields,...} = RecordsData.get thy;
+ val (flds,(more,_)) = split_last (Symtab.lookup_list extfields name);
+ val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
+
+ val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) (Vartab.empty);
+ val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds;
+ in (flds',(more,moreT)) end;
+
+fun get_recT_fields thy T =
+ let
+ val (root_flds,(root_more,root_moreT)) = get_extT_fields thy T;
+ val (rest_flds,rest_more) =
+ if is_recT root_moreT then get_recT_fields thy root_moreT
+ else ([],(root_more,root_moreT));
+ in (root_flds@rest_flds,rest_more) end;
+
+
+(* access 'fieldext' *)
+
+fun add_fieldext extname_types fields thy =
+ let
+ val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
+ RecordsData.get thy;
+ val fieldext' =
+ fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
+ val data=make_record_data records sel_upd equalities extinjects extsplit
+ splits extfields fieldext';
+ in RecordsData.put data thy end;
+
+
+val get_fieldext = Symtab.lookup o #fieldext o RecordsData.get;
+
+
+(* parent records *)
+
+fun add_parents thy NONE parents = parents
+ | add_parents thy (SOME (types, name)) parents =
+ let
+ fun err msg = error (msg ^ " parent record " ^ quote name);
+
+ val {args, parent, fields, extension, induct} =
+ (case get_record thy name of SOME info => info | NONE => err "Unknown");
+ val _ = if length types <> length args then err "Bad number of arguments for" else ();
+
+ fun bad_inst ((x, S), T) =
+ if Sign.of_sort thy (T, S) then NONE else SOME x
+ val bads = List.mapPartial bad_inst (args ~~ types);
+ val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
+
+ val inst = map fst args ~~ types;
+ val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
+ val parent' = Option.map (apfst (map subst)) parent;
+ val fields' = map (apsnd subst) fields;
+ val extension' = apsnd (map subst) extension;
+ in
+ add_parents thy parent'
+ (make_parent_info name fields' extension' induct :: parents)
+ end;
+
+
+
+(** concrete syntax for records **)
+
+(* decode type *)
+
+fun decode_type thy t =
+ let
+ fun get_sort xs n = AList.lookup (op =) xs (n: indexname) |> the_default (Sign.defaultS thy);
+ val map_sort = Sign.intern_sort thy;
+ in
+ Syntax.typ_of_term (get_sort (Syntax.term_sorts map_sort t)) map_sort t
+ |> Sign.intern_tycons thy
+ end;
+
+
+(* parse translations *)
+
+fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) =
+ if c = mark then Syntax.const (suffix sfx name) $ (Abs ("_",dummyT, arg))
+ else raise TERM ("gen_field_tr: " ^ mark, [t])
+ | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
+
+fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) =
+ if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u
+ else [gen_field_tr mark sfx tm]
+ | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm];
+
+
+fun record_update_tr [t, u] =
+ Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t)
+ | record_update_tr ts = raise TERM ("record_update_tr", ts);
+
+fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts
+ | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts
+ | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
+ (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts
+ | update_name_tr ts = raise TERM ("update_name_tr", ts);
+
+fun dest_ext_field mark (t as (Const (c,_) $ Const (name,_) $ arg)) =
+ if c = mark then (name,arg) else raise TERM ("dest_ext_field: " ^ mark, [t])
+ | dest_ext_field _ t = raise TERM ("dest_ext_field", [t])
+
+fun dest_ext_fields sep mark (trm as (Const (c,_) $ t $ u)) =
+ if c = sep then dest_ext_field mark t::dest_ext_fields sep mark u
+ else [dest_ext_field mark trm]
+ | dest_ext_fields _ mark t = [dest_ext_field mark t]
+
+fun gen_ext_fields_tr sep mark sfx more ctxt t =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val msg = "error in record input: ";
+ val fieldargs = dest_ext_fields sep mark t;
+ fun splitargs (field::fields) ((name,arg)::fargs) =
+ if can (unsuffix name) field
+ then let val (args,rest) = splitargs fields fargs
+ in (arg::args,rest) end
+ else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
+ | splitargs [] (fargs as (_::_)) = ([],fargs)
+ | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t])
+ | splitargs _ _ = ([],[]);
+
+ fun mk_ext (fargs as (name,arg)::_) =
+ (case get_fieldext thy (Sign.intern_const thy name) of
+ SOME (ext,_) => (case get_extfields thy ext of
+ SOME flds
+ => let val (args,rest) =
+ splitargs (map fst (but_last flds)) fargs;
+ val more' = mk_ext rest;
+ in list_comb (Syntax.const (suffix sfx ext),args@[more'])
+ end
+ | NONE => raise TERM(msg ^ "no fields defined for "
+ ^ ext,[t]))
+ | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
+ | mk_ext [] = more
+
+ in mk_ext fieldargs end;
+
+fun gen_ext_type_tr sep mark sfx more ctxt t =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val msg = "error in record-type input: ";
+ val fieldargs = dest_ext_fields sep mark t;
+ fun splitargs (field::fields) ((name,arg)::fargs) =
+ if can (unsuffix name) field
+ then let val (args,rest) = splitargs fields fargs
+ in (arg::args,rest) end
+ else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
+ | splitargs [] (fargs as (_::_)) = ([],fargs)
+ | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t])
+ | splitargs _ _ = ([],[]);
+
+ fun mk_ext (fargs as (name,arg)::_) =
+ (case get_fieldext thy (Sign.intern_const thy name) of
+ SOME (ext,alphas) =>
+ (case get_extfields thy ext of
+ SOME flds
+ => (let
+ val flds' = but_last flds;
+ val types = map snd flds';
+ val (args,rest) = splitargs (map fst flds') fargs;
+ val argtypes = map (Sign.certify_typ thy o decode_type thy) args;
+ val midx = fold (fn T => fn i => Int.max (maxidx_of_typ T, i))
+ argtypes 0;
+ val varifyT = varifyT midx;
+ val vartypes = map varifyT types;
+
+ val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes)
+ Vartab.empty;
+ val alphas' = map ((Syntax.term_of_typ (! Syntax.show_sorts)) o
+ Envir.norm_type subst o varifyT)
+ (but_last alphas);
+
+ val more' = mk_ext rest;
+ in list_comb (Syntax.const (suffix sfx ext),alphas'@[more'])
+ end handle TYPE_MATCH => raise
+ TERM (msg ^ "type is no proper record (extension)", [t]))
+ | NONE => raise TERM (msg ^ "no fields defined for " ^ ext,[t]))
+ | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
+ | mk_ext [] = more
+
+ in mk_ext fieldargs end;
+
+fun gen_adv_record_tr sep mark sfx unit ctxt [t] =
+ gen_ext_fields_tr sep mark sfx unit ctxt t
+ | gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
+
+fun gen_adv_record_scheme_tr sep mark sfx ctxt [t, more] =
+ gen_ext_fields_tr sep mark sfx more ctxt t
+ | gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
+
+fun gen_adv_record_type_tr sep mark sfx unit ctxt [t] =
+ gen_ext_type_tr sep mark sfx unit ctxt t
+ | gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
+
+fun gen_adv_record_type_scheme_tr sep mark sfx ctxt [t, more] =
+ gen_ext_type_tr sep mark sfx more ctxt t
+ | gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
+
+val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit;
+val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN;
+
+val adv_record_type_tr =
+ gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN
+ (Syntax.term_of_typ false (HOLogic.unitT));
+val adv_record_type_scheme_tr =
+ gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN;
+
+
+val parse_translation =
+ [("_record_update", record_update_tr),
+ ("_update_name", update_name_tr)];
+
+
+val adv_parse_translation =
+ [("_record",adv_record_tr),
+ ("_record_scheme",adv_record_scheme_tr),
+ ("_record_type",adv_record_type_tr),
+ ("_record_type_scheme",adv_record_type_scheme_tr)];
+
+
+(* print translations *)
+
+val print_record_type_abbr = ref true;
+val print_record_type_as_fields = ref true;
+
+fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ k $ u) =
+ let val t = (case k of (Abs (_,_,(Abs (_,_,t)$Bound 0)))
+ => if null (loose_bnos t) then t else raise Match
+ | Abs (x,_,t) => if null (loose_bnos t) then t else raise Match
+ | _ => raise Match)
+
+ (* (case k of (Const ("K_record",_)$t) => t
+ | Abs (x,_,Const ("K_record",_)$t$Bound 0) => t
+ | _ => raise Match)*)
+ in
+ (case try (unsuffix sfx) name_field of
+ SOME name =>
+ apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u)
+ | NONE => ([], tm))
+ end
+ | gen_field_upds_tr' _ _ tm = ([], tm);
+
+fun record_update_tr' tm =
+ let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in
+ if null ts then raise Match
+ else Syntax.const "_record_update" $ u $
+ foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts)
+ end;
+
+fun gen_field_tr' sfx tr' name =
+ let val name_sfx = suffix sfx name
+ in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end;
+
+fun record_tr' sep mark record record_scheme unit ctxt t =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ fun field_lst t =
+ (case strip_comb t of
+ (Const (ext,_),args as (_::_))
+ => (case try (unsuffix extN) (Sign.intern_const thy ext) of
+ SOME ext'
+ => (case get_extfields thy ext' of
+ SOME flds
+ => (let
+ val (f::fs) = but_last (map fst flds);
+ val flds' = Sign.extern_const thy f :: map Long_Name.base_name fs;
+ val (args',more) = split_last args;
+ in (flds'~~args')@field_lst more end
+ handle Library.UnequalLengths => [("",t)])
+ | NONE => [("",t)])
+ | NONE => [("",t)])
+ | _ => [("",t)])
+
+ val (flds,(_,more)) = split_last (field_lst t);
+ val _ = if null flds then raise Match else ();
+ val flds' = map (fn (n,t)=>Syntax.const mark$Syntax.const n$t) flds;
+ val flds'' = foldr1 (fn (x,y) => Syntax.const sep$x$y) flds';
+
+ in if unit more
+ then Syntax.const record$flds''
+ else Syntax.const record_scheme$flds''$more
+ end
+
+fun gen_record_tr' name =
+ let val name_sfx = suffix extN name;
+ val unit = (fn Const (@{const_syntax "Product_Type.Unity"},_) => true | _ => false);
+ fun tr' ctxt ts = record_tr' "_fields" "_field" "_record" "_record_scheme" unit ctxt
+ (list_comb (Syntax.const name_sfx,ts))
+ in (name_sfx,tr')
+ end
+
+fun print_translation names =
+ map (gen_field_tr' updateN record_update_tr') names;
+
+
+(* record_type_abbr_tr' tries to reconstruct the record name type abbreviation from *)
+(* the (nested) extension types. *)
+fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt tm =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ (* tm is term representation of a (nested) field type. We first reconstruct the *)
+ (* type from tm so that we can continue on the type level rather then the term level.*)
+
+ (* WORKAROUND:
+ * If a record type occurs in an error message of type inference there
+ * may be some internal frees donoted by ??:
+ * (Const "_tfree",_)$Free ("??'a",_).
+
+ * This will unfortunately be translated to Type ("??'a",[]) instead of
+ * TFree ("??'a",_) by typ_of_term, which will confuse unify below.
+ * fixT works around.
+ *)
+ fun fixT (T as Type (x,[])) =
+ if String.isPrefix "??'" x then TFree (x,Sign.defaultS thy) else T
+ | fixT (Type (x,xs)) = Type (x,map fixT xs)
+ | fixT T = T;
+
+ val T = fixT (decode_type thy tm);
+ val midx = maxidx_of_typ T;
+ val varifyT = varifyT midx;
+
+ fun mk_type_abbr subst name alphas =
+ let val abbrT = Type (name, map (fn a => varifyT (TFree (a, Sign.defaultS thy))) alphas);
+ in Syntax.term_of_typ (! Syntax.show_sorts)
+ (Sign.extern_typ thy (Envir.norm_type subst abbrT)) end;
+
+ fun match rT T = (Sign.typ_match thy (varifyT rT,T)
+ Vartab.empty);
+
+ in if !print_record_type_abbr
+ then (case last_extT T of
+ SOME (name,_)
+ => if name = lastExt
+ then
+ (let
+ val subst = match schemeT T
+ in
+ if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree(zeta,Sign.defaultS thy))))
+ then mk_type_abbr subst abbr alphas
+ else mk_type_abbr subst (suffix schemeN abbr) (alphas@[zeta])
+ end handle TYPE_MATCH => default_tr' ctxt tm)
+ else raise Match (* give print translation of specialised record a chance *)
+ | _ => raise Match)
+ else default_tr' ctxt tm
+ end
+
+fun record_type_tr' sep mark record record_scheme ctxt t =
+ let
+ val thy = ProofContext.theory_of ctxt;
+
+ val T = decode_type thy t;
+ val varifyT = varifyT (Term.maxidx_of_typ T);
+
+ fun term_of_type T = Syntax.term_of_typ (!Syntax.show_sorts) (Sign.extern_typ thy T);
+
+ fun field_lst T =
+ (case T of
+ Type (ext, args)
+ => (case try (unsuffix ext_typeN) ext of
+ SOME ext'
+ => (case get_extfields thy ext' of
+ SOME flds
+ => (case get_fieldext thy (fst (hd flds)) of
+ SOME (_, alphas)
+ => (let
+ val (f :: fs) = but_last flds;
+ val flds' = apfst (Sign.extern_const thy) f
+ :: map (apfst Long_Name.base_name) fs;
+ val (args', more) = split_last args;
+ val alphavars = map varifyT (but_last alphas);
+ val subst = fold2 (curry (Sign.typ_match thy))
+ alphavars args' Vartab.empty;
+ val flds'' = (map o apsnd)
+ (Envir.norm_type subst o varifyT) flds';
+ in flds'' @ field_lst more end
+ handle TYPE_MATCH => [("", T)]
+ | Library.UnequalLengths => [("", T)])
+ | NONE => [("", T)])
+ | NONE => [("", T)])
+ | NONE => [("", T)])
+ | _ => [("", T)])
+
+ val (flds, (_, moreT)) = split_last (field_lst T);
+ val flds' = map (fn (n, T) => Syntax.const mark $ Syntax.const n $ term_of_type T) flds;
+ val flds'' = foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds' handle Empty => raise Match;
+
+ in if not (!print_record_type_as_fields) orelse null flds then raise Match
+ else if moreT = HOLogic.unitT
+ then Syntax.const record$flds''
+ else Syntax.const record_scheme$flds''$term_of_type moreT
+ end
+
+
+fun gen_record_type_tr' name =
+ let val name_sfx = suffix ext_typeN name;
+ fun tr' ctxt ts = record_type_tr' "_field_types" "_field_type"
+ "_record_type" "_record_type_scheme" ctxt
+ (list_comb (Syntax.const name_sfx,ts))
+ in (name_sfx,tr')
+ end
+
+
+fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name =
+ let val name_sfx = suffix ext_typeN name;
+ val default_tr' = record_type_tr' "_field_types" "_field_type"
+ "_record_type" "_record_type_scheme"
+ fun tr' ctxt ts =
+ record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt
+ (list_comb (Syntax.const name_sfx,ts))
+ in (name_sfx, tr') end;
+
+(** record simprocs **)
+
+val record_quick_and_dirty_sensitive = ref false;
+
+
+fun quick_and_dirty_prove stndrd thy asms prop tac =
+ if !record_quick_and_dirty_sensitive andalso !quick_and_dirty
+ then Goal.prove (ProofContext.init thy) [] []
+ (Logic.list_implies (map Logic.varify asms,Logic.varify prop))
+ (K (SkipProof.cheat_tac @{theory HOL}))
+ (* standard can take quite a while for large records, thats why
+ * we varify the proposition manually here.*)
+ else let val prf = Goal.prove (ProofContext.init thy) [] asms prop tac;
+ in if stndrd then standard prf else prf end;
+
+fun quick_and_dirty_prf noopt opt () =
+ if !record_quick_and_dirty_sensitive andalso !quick_and_dirty
+ then noopt ()
+ else opt ();
+
+local
+fun abstract_over_fun_app (Abs (f,fT,t)) =
+ let
+ val (f',t') = Term.dest_abs (f,fT,t);
+ val T = domain_type fT;
+ val (x,T') = hd (Term.variant_frees t' [("x",T)]);
+ val f_x = Free (f',fT)$(Free (x,T'));
+ fun is_constr (Const (c,_)$_) = can (unsuffix extN) c
+ | is_constr _ = false;
+ fun subst (t as u$w) = if Free (f',fT)=u
+ then if is_constr w then f_x
+ else raise TERM ("abstract_over_fun_app",[t])
+ else subst u$subst w
+ | subst (Abs (x,T,t)) = (Abs (x,T,subst t))
+ | subst t = t
+ val t'' = abstract_over (f_x,subst t');
+ val vars = strip_qnt_vars "all" t'';
+ val bdy = strip_qnt_body "all" t'';
+
+ in list_abs ((x,T')::vars,bdy) end
+ | abstract_over_fun_app t = raise TERM ("abstract_over_fun_app",[t]);
+(* Generates a theorem of the kind:
+ * !!f x*. PROP P (f ( r x* ) x* == !!r x*. PROP P r x*
+ *)
+fun mk_fun_apply_eq (Abs (f, fT, t)) thy =
+ let
+ val rT = domain_type fT;
+ val vars = Term.strip_qnt_vars "all" t;
+ val Ts = map snd vars;
+ val n = length vars;
+ fun app_bounds 0 t = t$Bound 0
+ | app_bounds n t = if n > 0 then app_bounds (n-1) (t$Bound n) else t
+
+
+ val [P,r] = Term.variant_frees t [("P",rT::Ts--->Term.propT),("r",Ts--->rT)];
+ val prop = Logic.mk_equals
+ (list_all ((f,fT)::vars,
+ app_bounds (n - 1) ((Free P)$(Bound n$app_bounds (n-1) (Free r)))),
+ list_all ((fst r,rT)::vars,
+ app_bounds (n - 1) ((Free P)$Bound n)));
+ val prove_standard = quick_and_dirty_prove true thy;
+ val thm = prove_standard [] prop (fn _ =>
+ EVERY [rtac equal_intr_rule 1,
+ Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1,
+ Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1]);
+ in thm end
+ | mk_fun_apply_eq t thy = raise TERM ("mk_fun_apply_eq",[t]);
+
+in
+(* During proof of theorems produced by record_simproc you can end up in
+ * situations like "!!f ... . ... f r ..." where f is an extension update function.
+ * In order to split "f r" we transform this to "!!r ... . ... r ..." so that the
+ * usual split rules for extensions can apply.
+ *)
+val record_split_f_more_simproc =
+ Simplifier.simproc @{theory HOL} "record_split_f_more_simp" ["x"]
+ (fn thy => fn _ => fn t =>
+ (case t of (Const ("all", Type (_, [Type (_, [Type("fun",[T,T']), _]), _])))$
+ (trm as Abs _) =>
+ (case rec_id (~1) T of
+ "" => NONE
+ | n => if T=T'
+ then (let
+ val P=cterm_of thy (abstract_over_fun_app trm);
+ val thm = mk_fun_apply_eq trm thy;
+ val PV = cterm_of thy (hd (OldTerm.term_vars (prop_of thm)));
+ val thm' = cterm_instantiate [(PV,P)] thm;
+ in SOME thm' end handle TERM _ => NONE)
+ else NONE)
+ | _ => NONE))
+end
+
+fun prove_split_simp thy ss T prop =
+ let
+ val {sel_upd={simpset,...},extsplit,...} = RecordsData.get thy;
+ val extsplits =
+ Library.foldl (fn (thms,(n,_)) => the_list (Symtab.lookup extsplit n) @ thms)
+ ([],dest_recTs T);
+ val thms = (case get_splits thy (rec_id (~1) T) of
+ SOME (all_thm,_,_,_) =>
+ all_thm::(case extsplits of [thm] => [] | _ => extsplits)
+ (* [thm] is the same as all_thm *)
+ | NONE => extsplits)
+ val thms'=K_comp_convs@thms;
+ val ss' = (Simplifier.inherit_context ss simpset
+ addsimps thms'
+ addsimprocs [record_split_f_more_simproc]);
+ in
+ quick_and_dirty_prove true thy [] prop (fn _ => simp_tac ss' 1)
+ end;
+
+
+local
+fun eq (s1:string) (s2:string) = (s1 = s2);
+fun has_field extfields f T =
+ exists (fn (eN,_) => exists (eq f o fst) (Symtab.lookup_list extfields eN))
+ (dest_recTs T);
+
+fun K_skeleton n (T as Type (_,[_,kT])) (b as Bound i) (Abs (x,xT,t)) =
+ if null (loose_bnos t) then ((n,kT),(Abs (x,xT,Bound (i+1)))) else ((n,T),b)
+ | K_skeleton n T b _ = ((n,T),b);
+
+(*
+fun K_skeleton n _ b ((K_rec as Const ("Record.K_record",Type (_,[kT,_])))$_) =
+ ((n,kT),K_rec$b)
+ | K_skeleton n _ (Bound i)
+ (Abs (x,T,(K_rec as Const ("Record.K_record",Type (_,[kT,_])))$_$Bound 0)) =
+ ((n,kT),Abs (x,T,(K_rec$Bound (i+1)$Bound 0)))
+ | K_skeleton n T b _ = ((n,T),b);
+ *)
+
+fun normalize_rhs thm =
+ let
+ val ss = HOL_basic_ss addsimps K_comp_convs;
+ val rhs = thm |> Thm.cprop_of |> Thm.dest_comb |> snd;
+ val rhs' = (Simplifier.rewrite ss rhs);
+ in Thm.transitive thm rhs' end;
+in
+(* record_simproc *)
+(* Simplifies selections of an record update:
+ * (1) S (S_update k r) = k (S r)
+ * (2) S (X_update k r) = S r
+ * The simproc skips multiple updates at once, eg:
+ * S (X_update x (Y_update y (S_update k r))) = k (S r)
+ * But be careful in (2) because of the extendibility of records.
+ * - If S is a more-selector we have to make sure that the update on component
+ * X does not affect the selected subrecord.
+ * - If X is a more-selector we have to make sure that S is not in the updated
+ * subrecord.
+ *)
+val record_simproc =
+ Simplifier.simproc @{theory HOL} "record_simp" ["x"]
+ (fn thy => fn ss => fn t =>
+ (case t of (sel as Const (s, Type (_,[domS,rangeS])))$
+ ((upd as Const (u,Type(_,[_,Type (_,[rT,_])]))) $ k $ r)=>
+ if is_selector thy s then
+ (case get_updates thy u of SOME u_name =>
+ let
+ val {sel_upd={updates,...},extfields,...} = RecordsData.get thy;
+
+ fun mk_eq_terms ((upd as Const (u,Type(_,[kT,_]))) $ k $ r) =
+ (case Symtab.lookup updates u of
+ NONE => NONE
+ | SOME u_name
+ => if u_name = s
+ then (case mk_eq_terms r of
+ NONE =>
+ let
+ val rv = ("r",rT)
+ val rb = Bound 0
+ val (kv,kb) = K_skeleton "k" kT (Bound 1) k;
+ in SOME (upd$kb$rb,kb$(sel$rb),[kv,rv]) end
+ | SOME (trm,trm',vars) =>
+ let
+ val (kv,kb) = K_skeleton "k" kT (Bound (length vars)) k;
+ in SOME (upd$kb$trm,kb$trm',kv::vars) end)
+ else if has_field extfields u_name rangeS
+ orelse has_field extfields s (domain_type kT)
+ then NONE
+ else (case mk_eq_terms r of
+ SOME (trm,trm',vars)
+ => let
+ val (kv,kb) =
+ K_skeleton "k" kT (Bound (length vars)) k;
+ in SOME (upd$kb$trm,trm',kv::vars) end
+ | NONE
+ => let
+ val rv = ("r",rT)
+ val rb = Bound 0
+ val (kv,kb) = K_skeleton "k" kT (Bound 1) k;
+ in SOME (upd$kb$rb,sel$rb,[kv,rv]) end))
+ | mk_eq_terms r = NONE
+ in
+ (case mk_eq_terms (upd$k$r) of
+ SOME (trm,trm',vars)
+ => SOME (prove_split_simp thy ss domS
+ (list_all(vars, Logic.mk_equals (sel $ trm, trm'))))
+ | NONE => NONE)
+ end
+ | NONE => NONE)
+ else NONE
+ | _ => NONE));
+
+(* record_upd_simproc *)
+(* simplify multiple updates:
+ * (1) "N_update y (M_update g (N_update x (M_update f r))) =
+ (N_update (y o x) (M_update (g o f) r))"
+ * (2) "r(|M:= M r|) = r"
+ * For (2) special care of "more" updates has to be taken:
+ * r(|more := m; A := A r|)
+ * If A is contained in the fields of m we cannot remove the update A := A r!
+ * (But r(|more := r; A := A (r(|more := r|))|) = r(|more := r|)
+*)
+val record_upd_simproc =
+ Simplifier.simproc @{theory HOL} "record_upd_simp" ["x"]
+ (fn thy => fn ss => fn t =>
+ (case t of ((upd as Const (u, Type(_,[_,Type(_,[rT,_])]))) $ k $ r) =>
+ let datatype ('a,'b) calc = Init of 'b | Inter of 'a
+ val {sel_upd={selectors,updates,...},extfields,...} = RecordsData.get thy;
+
+ (*fun mk_abs_var x t = (x, fastype_of t);*)
+ fun sel_name u = Long_Name.base_name (unsuffix updateN u);
+
+ fun seed s (upd as Const (more,Type(_,[mT,_]))$ k $ r) =
+ if has_field extfields s (domain_type' mT) then upd else seed s r
+ | seed _ r = r;
+
+ fun grow u uT k kT vars (sprout,skeleton) =
+ if sel_name u = moreN
+ then let val (kv,kb) = K_skeleton "k" kT (Bound (length vars)) k;
+ in ((Const (u,uT)$k$sprout,Const (u,uT)$kb$skeleton),kv::vars) end
+ else ((sprout,skeleton),vars);
+
+
+ fun dest_k (Abs (x,T,((sel as Const (s,_))$r))) =
+ if null (loose_bnos r) then SOME (x,T,sel,s,r) else NONE
+ | dest_k (Abs (_,_,(Abs (x,T,((sel as Const (s,_))$r)))$Bound 0)) =
+ (* eta expanded variant *)
+ if null (loose_bnos r) then SOME (x,T,sel,s,r) else NONE
+ | dest_k _ = NONE;
+
+ fun is_upd_same (sprout,skeleton) u k =
+ (case dest_k k of SOME (x,T,sel,s,r) =>
+ if (unsuffix updateN u) = s andalso (seed s sprout) = r
+ then SOME (fn t => Abs (x,T,incr_boundvars 1 t),sel,seed s skeleton)
+ else NONE
+ | NONE => NONE);
+
+ fun init_seed r = ((r,Bound 0), [("r", rT)]);
+
+ fun add (n:string) f fmaps =
+ (case AList.lookup (op =) fmaps n of
+ NONE => AList.update (op =) (n,[f]) fmaps
+ | SOME fs => AList.update (op =) (n,f::fs) fmaps)
+
+ fun comps (n:string) T fmaps =
+ (case AList.lookup (op =) fmaps n of
+ SOME fs =>
+ foldr1 (fn (f,g) => Const ("Fun.comp",(T-->T)-->(T-->T)-->(T-->T))$f$g) fs
+ | NONE => error ("record_upd_simproc.comps"))
+
+ (* mk_updterm returns either
+ * - Init (orig-term, orig-term-skeleton, vars) if no optimisation can be made,
+ * where vars are the bound variables in the skeleton
+ * - Inter (orig-term-skeleton,simplified-term-skeleton,
+ * vars, (term-sprout, skeleton-sprout))
+ * where "All vars. orig-term-skeleton = simplified-term-skeleton" is
+ * the desired simplification rule,
+ * the sprouts accumulate the "more-updates" on the way from the seed
+ * to the outermost update. It is only relevant to calculate the
+ * possible simplification for (2)
+ * The algorithm first walks down the updates to the seed-record while
+ * memorising the updates in the already-table. While walking up the
+ * updates again, the optimised term is constructed.
+ *)
+ fun mk_updterm upds already
+ (t as ((upd as Const (u,uT as (Type (_,[kT,_])))) $ k $ r)) =
+ if Symtab.defined upds u
+ then let
+ fun rest already = mk_updterm upds already
+ in if u mem_string already
+ then (case (rest already r) of
+ Init ((sprout,skel),vars) =>
+ let
+ val n = sel_name u;
+ val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
+ val (sprout',vars')= grow u uT k kT (kv::vars) (sprout,skel);
+ in Inter (upd$kb$skel,skel,vars',add n kb [],sprout') end
+ | Inter (trm,trm',vars,fmaps,sprout) =>
+ let
+ val n = sel_name u;
+ val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
+ val (sprout',vars') = grow u uT k kT (kv::vars) sprout;
+ in Inter(upd$kb$trm,trm',kv::vars',add n kb fmaps,sprout')
+ end)
+ else
+ (case rest (u::already) r of
+ Init ((sprout,skel),vars) =>
+ (case is_upd_same (sprout,skel) u k of
+ SOME (K_rec,sel,skel') =>
+ let
+ val (sprout',vars') = grow u uT k kT vars (sprout,skel);
+ in Inter(upd$(K_rec (sel$skel'))$skel,skel,vars',[],sprout')
+ end
+ | NONE =>
+ let
+ val n = sel_name u;
+ val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
+ in Init ((upd$k$sprout,upd$kb$skel),kv::vars) end)
+ | Inter (trm,trm',vars,fmaps,sprout) =>
+ (case is_upd_same sprout u k of
+ SOME (K_rec,sel,skel) =>
+ let
+ val (sprout',vars') = grow u uT k kT vars sprout
+ in Inter(upd$(K_rec (sel$skel))$trm,trm',vars',fmaps,sprout')
+ end
+ | NONE =>
+ let
+ val n = sel_name u
+ val T = domain_type kT
+ val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
+ val (sprout',vars') = grow u uT k kT (kv::vars) sprout
+ val fmaps' = add n kb fmaps
+ in Inter (upd$kb$trm,upd$comps n T fmaps'$trm'
+ ,vars',fmaps',sprout') end))
+ end
+ else Init (init_seed t)
+ | mk_updterm _ _ t = Init (init_seed t);
+
+ in (case mk_updterm updates [] t of
+ Inter (trm,trm',vars,_,_)
+ => SOME (normalize_rhs
+ (prove_split_simp thy ss rT
+ (list_all(vars, Logic.mk_equals (trm, trm')))))
+ | _ => NONE)
+ end
+ | _ => NONE))
+end
+
+(* record_eq_simproc *)
+(* looks up the most specific record-equality.
+ * Note on efficiency:
+ * Testing equality of records boils down to the test of equality of all components.
+ * Therefore the complexity is: #components * complexity for single component.
+ * Especially if a record has a lot of components it may be better to split up
+ * the record first and do simplification on that (record_split_simp_tac).
+ * e.g. r(|lots of updates|) = x
+ *
+ * record_eq_simproc record_split_simp_tac
+ * Complexity: #components * #updates #updates
+ *
+ *)
+val record_eq_simproc =
+ Simplifier.simproc @{theory HOL} "record_eq_simp" ["r = s"]
+ (fn thy => fn _ => fn t =>
+ (case t of Const ("op =", Type (_, [T, _])) $ _ $ _ =>
+ (case rec_id (~1) T of
+ "" => NONE
+ | name => (case get_equalities thy name of
+ NONE => NONE
+ | SOME thm => SOME (thm RS Eq_TrueI)))
+ | _ => NONE));
+
+(* record_split_simproc *)
+(* splits quantified occurrences of records, for which P holds. P can peek on the
+ * subterm starting at the quantified occurrence of the record (including the quantifier)
+ * P t = 0: do not split
+ * P t = ~1: completely split
+ * P t > 0: split up to given bound of record extensions
+ *)
+fun record_split_simproc P =
+ Simplifier.simproc @{theory HOL} "record_split_simp" ["x"]
+ (fn thy => fn _ => fn t =>
+ (case t of (Const (quantifier, Type (_, [Type (_, [T, _]), _])))$trm =>
+ if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex"
+ then (case rec_id (~1) T of
+ "" => NONE
+ | name
+ => let val split = P t
+ in if split <> 0 then
+ (case get_splits thy (rec_id split T) of
+ NONE => NONE
+ | SOME (all_thm, All_thm, Ex_thm,_)
+ => SOME (case quantifier of
+ "all" => all_thm
+ | "All" => All_thm RS eq_reflection
+ | "Ex" => Ex_thm RS eq_reflection
+ | _ => error "record_split_simproc"))
+ else NONE
+ end)
+ else NONE
+ | _ => NONE))
+
+val record_ex_sel_eq_simproc =
+ Simplifier.simproc @{theory HOL} "record_ex_sel_eq_simproc" ["Ex t"]
+ (fn thy => fn ss => fn t =>
+ let
+ fun prove prop =
+ quick_and_dirty_prove true thy [] prop
+ (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy)
+ addsimps simp_thms addsimprocs [record_split_simproc (K ~1)]) 1);
+
+ fun mkeq (lr,Teq,(sel,Tsel),x) i =
+ if is_selector thy sel then
+ let val x' = if not (loose_bvar1 (x,0))
+ then Free ("x" ^ string_of_int i, range_type Tsel)
+ else raise TERM ("",[x]);
+ val sel' = Const (sel,Tsel)$Bound 0;
+ val (l,r) = if lr then (sel',x') else (x',sel');
+ in Const ("op =",Teq)$l$r end
+ else raise TERM ("",[Const (sel,Tsel)]);
+
+ fun dest_sel_eq (Const ("op =",Teq)$(Const (sel,Tsel)$Bound 0)$X) =
+ (true,Teq,(sel,Tsel),X)
+ | dest_sel_eq (Const ("op =",Teq)$X$(Const (sel,Tsel)$Bound 0)) =
+ (false,Teq,(sel,Tsel),X)
+ | dest_sel_eq _ = raise TERM ("",[]);
+
+ in
+ (case t of
+ (Const ("Ex",Tex)$Abs(s,T,t)) =>
+ (let val eq = mkeq (dest_sel_eq t) 0;
+ val prop = list_all ([("r",T)],
+ Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
+ HOLogic.true_const));
+ in SOME (prove prop) end
+ handle TERM _ => NONE)
+ | _ => NONE)
+ end)
+
+
+
+
+local
+val inductive_atomize = thms "induct_atomize";
+val inductive_rulify = thms "induct_rulify";
+in
+(* record_split_simp_tac *)
+(* splits (and simplifies) all records in the goal for which P holds.
+ * For quantified occurrences of a record
+ * P can peek on the whole subterm (including the quantifier); for free variables P
+ * can only peek on the variable itself.
+ * P t = 0: do not split
+ * P t = ~1: completely split
+ * P t > 0: split up to given bound of record extensions
+ *)
+fun record_split_simp_tac thms P i st =
+ let
+ val thy = Thm.theory_of_thm st;
+
+ val has_rec = exists_Const
+ (fn (s, Type (_, [Type (_, [T, _]), _])) =>
+ (s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T
+ | _ => false);
+
+ val goal = nth (Thm.prems_of st) (i - 1);
+ val frees = List.filter (is_recT o type_of) (OldTerm.term_frees goal);
+
+ fun mk_split_free_tac free induct_thm i =
+ let val cfree = cterm_of thy free;
+ val (_$(_$r)) = concl_of induct_thm;
+ val crec = cterm_of thy r;
+ val thm = cterm_instantiate [(crec,cfree)] induct_thm;
+ in EVERY [simp_tac (HOL_basic_ss addsimps inductive_atomize) i,
+ rtac thm i,
+ simp_tac (HOL_basic_ss addsimps inductive_rulify) i]
+ end;
+
+ fun split_free_tac P i (free as Free (n,T)) =
+ (case rec_id (~1) T of
+ "" => NONE
+ | name => let val split = P free
+ in if split <> 0 then
+ (case get_splits thy (rec_id split T) of
+ NONE => NONE
+ | SOME (_,_,_,induct_thm)
+ => SOME (mk_split_free_tac free induct_thm i))
+ else NONE
+ end)
+ | split_free_tac _ _ _ = NONE;
+
+ val split_frees_tacs = List.mapPartial (split_free_tac P i) frees;
+
+ val simprocs = if has_rec goal then [record_split_simproc P] else [];
+ val thms' = K_comp_convs@thms
+ in st |> ((EVERY split_frees_tacs)
+ THEN (Simplifier.full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i))
+ end handle Empty => Seq.empty;
+end;
+
+
+(* record_split_tac *)
+(* splits all records in the goal, which are quantified by ! or !!. *)
+fun record_split_tac i st =
+ let
+ val thy = Thm.theory_of_thm st;
+
+ val has_rec = exists_Const
+ (fn (s, Type (_, [Type (_, [T, _]), _])) =>
+ (s = "all" orelse s = "All") andalso is_recT T
+ | _ => false);
+
+ val goal = nth (Thm.prems_of st) (i - 1);
+
+ fun is_all t =
+ (case t of (Const (quantifier, _)$_) =>
+ if quantifier = "All" orelse quantifier = "all" then ~1 else 0
+ | _ => 0);
+
+ in if has_rec goal
+ then Simplifier.full_simp_tac
+ (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i st
+ else Seq.empty
+ end handle Subscript => Seq.empty;
+
+(* wrapper *)
+
+val record_split_name = "record_split_tac";
+val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac);
+
+
+
+(** theory extender interface **)
+
+(* prepare arguments *)
+
+fun read_raw_parent ctxt raw_T =
+ (case ProofContext.read_typ_abbrev ctxt raw_T of
+ Type (name, Ts) => (Ts, name)
+ | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
+
+fun read_typ ctxt raw_T env =
+ let
+ val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;
+ val T = Syntax.read_typ ctxt' raw_T;
+ val env' = OldTerm.add_typ_tfrees (T, env);
+ in (T, env') end;
+
+fun cert_typ ctxt raw_T env =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg;
+ val env' = OldTerm.add_typ_tfrees (T, env);
+ in (T, env') end;
+
+
+(* attributes *)
+
+fun case_names_fields x = RuleCases.case_names ["fields"] x;
+fun induct_type_global name = [case_names_fields, Induct.induct_type name];
+fun cases_type_global name = [case_names_fields, Induct.cases_type name];
+
+(* tactics *)
+
+fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps));
+
+(* do case analysis / induction according to rule on last parameter of ith subgoal
+ * (or on s if there are no parameters);
+ * Instatiation of record variable (and predicate) in rule is calculated to
+ * avoid problems with higher order unification.
+ *)
+
+fun try_param_tac s rule i st =
+ let
+ val cert = cterm_of (Thm.theory_of_thm st);
+ val g = nth (prems_of st) (i - 1);
+ val params = Logic.strip_params g;
+ val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g);
+ val rule' = Thm.lift_rule (Thm.cprem_of st i) rule;
+ val (P, ys) = strip_comb (HOLogic.dest_Trueprop
+ (Logic.strip_assums_concl (prop_of rule')));
+ (* ca indicates if rule is a case analysis or induction rule *)
+ val (x, ca) = (case rev (Library.drop (length params, ys)) of
+ [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
+ (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
+ | [x] => (head_of x, false));
+ val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of
+ [] => (case AList.lookup (op =) (map dest_Free (OldTerm.term_frees (prop_of st))) s of
+ NONE => sys_error "try_param_tac: no such variable"
+ | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl),
+ (x, Free (s, T))])
+ | (_, T) :: _ => [(P, list_abs (params, if ca then concl
+ else incr_boundvars 1 (Abs (s, T, concl)))),
+ (x, list_abs (params, Bound 0))])) rule'
+ in compose_tac (false, rule'', nprems_of rule) i st end;
+
+
+(* !!x1 ... xn. ... ==> EX x1 ... xn. P x1 ... xn;
+ instantiates x1 ... xn with parameters x1 ... xn *)
+fun ex_inst_tac i st =
+ let
+ val thy = Thm.theory_of_thm st;
+ val g = nth (prems_of st) (i - 1);
+ val params = Logic.strip_params g;
+ val exI' = Thm.lift_rule (Thm.cprem_of st i) exI;
+ val (_$(_$x)) = Logic.strip_assums_concl (hd (prems_of exI'));
+ val cx = cterm_of thy (fst (strip_comb x));
+
+ in Seq.single (Library.foldl (fn (st,v) =>
+ Seq.hd
+ (compose_tac (false, cterm_instantiate
+ [(cx,cterm_of thy (list_abs (params,Bound v)))] exI',1)
+ i st)) (st,((length params) - 1) downto 0))
+ end;
+
+fun extension_typedef name repT alphas thy =
+ let
+ fun get_thms thy name =
+ let
+ val SOME { Abs_induct = abs_induct,
+ Abs_inject=abs_inject, Abs_inverse = abs_inverse, ...} = Typedef.get_info thy name;
+ val rewrite_rule = MetaSimplifier.rewrite_rule [rec_UNIV_I, rec_True_simp];
+ in map rewrite_rule [abs_inject, abs_inverse, abs_induct] end;
+ val tname = Binding.name (Long_Name.base_name name);
+ in
+ thy
+ |> Typecopy.typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
+ |-> (fn (name, _) => `(fn thy => get_thms thy name))
+ end;
+
+fun mixit convs refls =
+ let fun f ((res,lhs,rhs),refl) = ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs);
+ in #1 (Library.foldl f (([],[],convs),refls)) end;
+
+
+fun extension_definition full name fields names alphas zeta moreT more vars thy =
+ let
+ val base = Long_Name.base_name;
+ val fieldTs = (map snd fields);
+ val alphas_zeta = alphas@[zeta];
+ val alphas_zetaTs = map (fn n => TFree (n, HOLogic.typeS)) alphas_zeta;
+ val vT = TFree (Name.variant alphas_zeta "'v", HOLogic.typeS);
+ val extT_name = suffix ext_typeN name
+ val extT = Type (extT_name, alphas_zetaTs);
+ val repT = foldr1 HOLogic.mk_prodT (fieldTs@[moreT]);
+ val fields_more = fields@[(full moreN,moreT)];
+ val fields_moreTs = fieldTs@[moreT];
+ val bfields_more = map (apfst base) fields_more;
+ val r = Free (rN,extT)
+ val len = length fields;
+ val idxms = 0 upto len;
+
+ (* prepare declarations and definitions *)
+
+ (*fields constructor*)
+ val ext_decl = (mk_extC (name,extT) fields_moreTs);
+ (*
+ val ext_spec = Const ext_decl :==
+ (foldr (uncurry lambda)
+ (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more]))) (vars@[more]))
+ *)
+ val ext_spec = list_comb (Const ext_decl,vars@[more]) :==
+ (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more])));
+
+ fun mk_ext args = list_comb (Const ext_decl, args);
+
+ (*destructors*)
+ val _ = timing_msg "record extension preparing definitions";
+ val dest_decls = map (mk_selC extT o (apfst (suffix ext_dest))) bfields_more;
+
+ fun mk_dest_spec (i, (c,T)) =
+ let val snds = (funpow i HOLogic.mk_snd (mk_Rep name repT extT $ r))
+ in Const (mk_selC extT (suffix ext_dest c,T))
+ :== (lambda r (if i=len then snds else HOLogic.mk_fst snds))
+ end;
+ val dest_specs =
+ ListPair.map mk_dest_spec (idxms, fields_more);
+
+ (*updates*)
+ val upd_decls = map (mk_updC updN extT) bfields_more;
+ fun mk_upd_spec (c,T) =
+ let
+ val args = map (fn (n,nT) => if n=c then Free (base c,T --> T)$
+ (mk_sel r (suffix ext_dest n,nT))
+ else (mk_sel r (suffix ext_dest n,nT)))
+ fields_more;
+ in Const (mk_updC updN extT (c,T))$(Free (base c,T --> T))$r
+ :== mk_ext args
+ end;
+ val upd_specs = map mk_upd_spec fields_more;
+
+ (* 1st stage: defs_thy *)
+ fun mk_defs () =
+ thy
+ |> extension_typedef name repT (alphas @ [zeta])
+ ||> Sign.add_consts_i
+ (map (Syntax.no_syn o apfst Binding.name) (apfst base ext_decl :: dest_decls @ upd_decls))
+ ||>> PureThy.add_defs false
+ (map (Thm.no_attributes o apfst Binding.name) (ext_spec :: dest_specs))
+ ||>> PureThy.add_defs false
+ (map (Thm.no_attributes o apfst Binding.name) upd_specs)
+ |-> (fn args as ((_, dest_defs), upd_defs) =>
+ fold Code.add_default_eqn dest_defs
+ #> fold Code.add_default_eqn upd_defs
+ #> pair args);
+ val ((([abs_inject, abs_inverse, abs_induct], ext_def :: dest_defs), upd_defs), defs_thy) =
+ timeit_msg "record extension type/selector/update defs:" mk_defs;
+
+ (* prepare propositions *)
+ val _ = timing_msg "record extension preparing propositions";
+ val vars_more = vars@[more];
+ val named_vars_more = (names@[full moreN])~~vars_more;
+ val variants = map (fn (Free (x,_))=>x) vars_more;
+ val ext = mk_ext vars_more;
+ val s = Free (rN, extT);
+ val w = Free (wN, extT);
+ val P = Free (Name.variant variants "P", extT-->HOLogic.boolT);
+ val C = Free (Name.variant variants "C", HOLogic.boolT);
+
+ val inject_prop =
+ let val vars_more' = map (fn (Free (x,T)) => Free (x ^ "'",T)) vars_more;
+ in All (map dest_Free (vars_more@vars_more'))
+ ((HOLogic.eq_const extT $
+ mk_ext vars_more$mk_ext vars_more')
+ ===
+ foldr1 HOLogic.mk_conj (map HOLogic.mk_eq (vars_more ~~ vars_more')))
+ end;
+
+ val induct_prop =
+ (All (map dest_Free vars_more) (Trueprop (P $ ext)), Trueprop (P $ s));
+
+ val cases_prop =
+ (All (map dest_Free vars_more)
+ (Trueprop (HOLogic.mk_eq (s,ext)) ==> Trueprop C))
+ ==> Trueprop C;
+
+ (*destructors*)
+ val dest_conv_props =
+ map (fn (c, x as Free (_,T)) => mk_sel ext (suffix ext_dest c,T) === x) named_vars_more;
+
+ (*updates*)
+ fun mk_upd_prop (i,(c,T)) =
+ let val x' = Free (Name.variant variants (base c ^ "'"),T --> T)
+ val args' = nth_map i (K (x'$nth vars_more i)) vars_more
+ in mk_upd updN c x' ext === mk_ext args' end;
+ val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
+
+ val surjective_prop =
+ let val args =
+ map (fn (c, Free (_,T)) => mk_sel s (suffix ext_dest c,T)) named_vars_more;
+ in s === mk_ext args end;
+
+ val split_meta_prop =
+ let val P = Free (Name.variant variants "P", extT-->Term.propT) in
+ Logic.mk_equals
+ (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext))
+ end;
+
+ fun prove stndrd = quick_and_dirty_prove stndrd defs_thy;
+ val prove_standard = quick_and_dirty_prove true defs_thy;
+ fun prove_simp stndrd simps =
+ let val tac = simp_all_tac HOL_ss simps
+ in fn prop => prove stndrd [] prop (K tac) end;
+
+ fun inject_prf () = (prove_simp true [ext_def,abs_inject,Pair_eq] inject_prop);
+ val inject = timeit_msg "record extension inject proof:" inject_prf;
+
+ fun induct_prf () =
+ let val (assm, concl) = induct_prop
+ in prove_standard [assm] concl (fn {prems, ...} =>
+ EVERY [try_param_tac rN abs_induct 1,
+ simp_tac (HOL_ss addsimps [split_paired_all]) 1,
+ resolve_tac (map (rewrite_rule [ext_def]) prems) 1])
+ end;
+ val induct = timeit_msg "record extension induct proof:" induct_prf;
+
+ fun cases_prf_opt () =
+ let
+ val (_$(Pvar$_)) = concl_of induct;
+ val ind = cterm_instantiate
+ [(cterm_of defs_thy Pvar, cterm_of defs_thy
+ (lambda w (HOLogic.imp$HOLogic.mk_eq(r,w)$C)))]
+ induct;
+ in standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
+
+ fun cases_prf_noopt () =
+ prove_standard [] cases_prop (fn _ =>
+ EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1,
+ try_param_tac rN induct 1,
+ rtac impI 1,
+ REPEAT (etac allE 1),
+ etac mp 1,
+ rtac refl 1])
+
+ val cases_prf = quick_and_dirty_prf cases_prf_noopt cases_prf_opt;
+ val cases = timeit_msg "record extension cases proof:" cases_prf;
+
+ fun dest_convs_prf () = map (prove_simp false
+ ([ext_def,abs_inverse]@Pair_sel_convs@dest_defs)) dest_conv_props;
+ val dest_convs = timeit_msg "record extension dest_convs proof:" dest_convs_prf;
+ fun dest_convs_standard_prf () = map standard dest_convs;
+
+ val dest_convs_standard =
+ timeit_msg "record extension dest_convs_standard proof:" dest_convs_standard_prf;
+
+ fun upd_convs_prf_noopt () = map (prove_simp true (dest_convs_standard@upd_defs))
+ upd_conv_props;
+ fun upd_convs_prf_opt () =
+ let
+
+ fun mkrefl (c,T) = Thm.reflexive
+ (cterm_of defs_thy (Free (Name.variant variants (base c ^ "'"),T-->T)));
+ val refls = map mkrefl fields_more;
+ val dest_convs' = map mk_meta_eq dest_convs;
+ val map_eqs = map (uncurry Thm.combination) (refls ~~ dest_convs');
+
+ val constr_refl = Thm.reflexive (cterm_of defs_thy (head_of ext));
+
+ fun mkthm (udef,(fld_refl,thms)) =
+ let val bdyeq = Library.foldl (uncurry Thm.combination) (constr_refl,thms);
+ (* (|N=N (|N=N,M=M,K=K,more=more|)
+ M=M (|N=N,M=M,K=K,more=more|)
+ K=K'
+ more = more (|N=N,M=M,K=K,more=more|) =
+ (|N=N,M=M,K=K',more=more|)
+ *)
+ val (_$(_$v$r)$_) = prop_of udef;
+ val (_$(v'$_)$_) = prop_of fld_refl;
+ val udef' = cterm_instantiate
+ [(cterm_of defs_thy v,cterm_of defs_thy v'),
+ (cterm_of defs_thy r,cterm_of defs_thy ext)] udef;
+ in standard (Thm.transitive udef' bdyeq) end;
+ in map mkthm (rev upd_defs ~~ (mixit dest_convs' map_eqs)) end;
+
+ val upd_convs_prf = quick_and_dirty_prf upd_convs_prf_noopt upd_convs_prf_opt;
+
+ val upd_convs =
+ timeit_msg "record extension upd_convs proof:" upd_convs_prf;
+
+ fun surjective_prf () =
+ prove_standard [] surjective_prop (fn _ =>
+ (EVERY [try_param_tac rN induct 1,
+ simp_tac (HOL_basic_ss addsimps dest_convs_standard) 1]));
+ val surjective = timeit_msg "record extension surjective proof:" surjective_prf;
+
+ fun split_meta_prf () =
+ prove_standard [] split_meta_prop (fn _ =>
+ EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
+ etac meta_allE 1, atac 1,
+ rtac (prop_subst OF [surjective]) 1,
+ REPEAT (etac meta_allE 1), atac 1]);
+ val split_meta = timeit_msg "record extension split_meta proof:" split_meta_prf;
+
+
+ val (([inject',induct',cases',surjective',split_meta'],
+ [dest_convs',upd_convs']),
+ thm_thy) =
+ defs_thy
+ |> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
+ [("ext_inject", inject),
+ ("ext_induct", induct),
+ ("ext_cases", cases),
+ ("ext_surjective", surjective),
+ ("ext_split", split_meta)]
+ ||>> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
+ [("dest_convs", dest_convs_standard), ("upd_convs", upd_convs)]
+
+ in (thm_thy,extT,induct',inject',dest_convs',split_meta',upd_convs')
+ end;
+
+fun chunks [] [] = []
+ | chunks [] xs = [xs]
+ | chunks (l::ls) xs = Library.take (l,xs)::chunks ls (Library.drop (l,xs));
+
+fun chop_last [] = error "last: list should not be empty"
+ | chop_last [x] = ([],x)
+ | chop_last (x::xs) = let val (tl,l) = chop_last xs in (x::tl,l) end;
+
+fun subst_last s [] = error "subst_last: list should not be empty"
+ | subst_last s ([x]) = [s]
+ | subst_last s (x::xs) = (x::subst_last s xs);
+
+(* mk_recordT builds up the record type from the current extension tpye extT and a list
+ * of parent extensions, starting with the root of the record hierarchy
+*)
+fun mk_recordT extT =
+ fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT;
+
+
+
+fun obj_to_meta_all thm =
+ let
+ fun E thm = case (SOME (spec OF [thm]) handle THM _ => NONE) of
+ SOME thm' => E thm'
+ | NONE => thm;
+ val th1 = E thm;
+ val th2 = Drule.forall_intr_vars th1;
+ in th2 end;
+
+fun meta_to_obj_all thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val prop = Thm.prop_of thm;
+ val params = Logic.strip_params prop;
+ val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl prop);
+ val ct = cterm_of thy
+ (HOLogic.mk_Trueprop (HOLogic.list_all (params, concl)));
+ val thm' = Seq.hd (REPEAT (rtac allI 1) (Thm.trivial ct));
+ in
+ Thm.implies_elim thm' thm
+ end;
+
+
+
+(* record_definition *)
+
+fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy =
+ let
+ val external_names = NameSpace.external_names (Sign.naming_of thy);
+
+ val alphas = map fst args;
+ val name = Sign.full_bname thy bname;
+ val full = Sign.full_bname_path thy bname;
+ val base = Long_Name.base_name;
+
+ val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields);
+
+ val parent_fields = List.concat (map #fields parents);
+ val parent_chunks = map (length o #fields) parents;
+ val parent_names = map fst parent_fields;
+ val parent_types = map snd parent_fields;
+ val parent_fields_len = length parent_fields;
+ val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map base parent_names);
+ val parent_vars = ListPair.map Free (parent_variants, parent_types);
+ val parent_len = length parents;
+ val parents_idx = (map #name parents) ~~ (0 upto (parent_len - 1));
+
+ val fields = map (apfst full) bfields;
+ val names = map fst fields;
+ val extN = full bname;
+ val types = map snd fields;
+ val alphas_fields = List.foldr OldTerm.add_typ_tfree_names [] types;
+ val alphas_ext = alphas inter alphas_fields;
+ val len = length fields;
+ val variants =
+ Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) (map fst bfields);
+ val vars = ListPair.map Free (variants, types);
+ val named_vars = names ~~ vars;
+ val idxs = 0 upto (len - 1);
+ val idxms = 0 upto len;
+
+ val all_fields = parent_fields @ fields;
+ val all_names = parent_names @ names;
+ val all_types = parent_types @ types;
+ val all_len = parent_fields_len + len;
+ val all_variants = parent_variants @ variants;
+ val all_vars = parent_vars @ vars;
+ val all_named_vars = (parent_names ~~ parent_vars) @ named_vars;
+
+
+ val zeta = Name.variant alphas "'z";
+ val moreT = TFree (zeta, HOLogic.typeS);
+ val more = Free (moreN, moreT);
+ val full_moreN = full moreN;
+ val bfields_more = bfields @ [(moreN,moreT)];
+ val fields_more = fields @ [(full_moreN,moreT)];
+ val vars_more = vars @ [more];
+ val named_vars_more = named_vars @[(full_moreN,more)];
+ val all_vars_more = all_vars @ [more];
+ val all_named_vars_more = all_named_vars @ [(full_moreN,more)];
+
+ (* 1st stage: extension_thy *)
+ val (extension_thy,extT,ext_induct,ext_inject,ext_dest_convs,ext_split,u_convs) =
+ thy
+ |> Sign.add_path bname
+ |> extension_definition full extN fields names alphas_ext zeta moreT more vars;
+
+ val _ = timing_msg "record preparing definitions";
+ val Type extension_scheme = extT;
+ val extension_name = unsuffix ext_typeN (fst extension_scheme);
+ val extension = let val (n,Ts) = extension_scheme in (n,subst_last HOLogic.unitT Ts) end;
+ val extension_names =
+ (map ((unsuffix ext_typeN) o fst o #extension) parents) @ [extN];
+ val extension_id = Library.foldl (op ^) ("",extension_names);
+
+
+ fun rec_schemeT n = mk_recordT (map #extension (prune n parents)) extT;
+ val rec_schemeT0 = rec_schemeT 0;
+
+ fun recT n =
+ let val (c,Ts) = extension
+ in mk_recordT (map #extension (prune n parents)) (Type (c,subst_last HOLogic.unitT Ts))
+ end;
+ val recT0 = recT 0;
+
+ fun mk_rec args n =
+ let val (args',more) = chop_last args;
+ fun mk_ext' (((name,T),args),more) = mk_ext (name,T) (args@[more]);
+ fun build Ts =
+ List.foldr mk_ext' more (prune n (extension_names ~~ Ts ~~ (chunks parent_chunks args')))
+ in
+ if more = HOLogic.unit
+ then build (map recT (0 upto parent_len))
+ else build (map rec_schemeT (0 upto parent_len))
+ end;
+
+ val r_rec0 = mk_rec all_vars_more 0;
+ val r_rec_unit0 = mk_rec (all_vars@[HOLogic.unit]) 0;
+
+ fun r n = Free (rN, rec_schemeT n)
+ val r0 = r 0;
+ fun r_unit n = Free (rN, recT n)
+ val r_unit0 = r_unit 0;
+ val w = Free (wN, rec_schemeT 0)
+
+ (* prepare print translation functions *)
+ val field_tr's =
+ print_translation (distinct (op =) (maps external_names (full_moreN :: names)));
+
+ val adv_ext_tr's =
+ let
+ val trnames = external_names extN;
+ in map (gen_record_tr') trnames end;
+
+ val adv_record_type_abbr_tr's =
+ let val trnames = external_names (hd extension_names);
+ val lastExt = unsuffix ext_typeN (fst extension);
+ in map (gen_record_type_abbr_tr' name alphas zeta lastExt rec_schemeT0) trnames
+ end;
+
+ val adv_record_type_tr's =
+ let val trnames = if parent_len > 0 then external_names extN else [];
+ (* avoid conflict with adv_record_type_abbr_tr's *)
+ in map (gen_record_type_tr') trnames
+ end;
+
+
+ (* prepare declarations *)
+
+ val sel_decls = map (mk_selC rec_schemeT0) bfields_more;
+ val upd_decls = map (mk_updC updateN rec_schemeT0) bfields_more;
+ val make_decl = (makeN, all_types ---> recT0);
+ val fields_decl = (fields_selN, types ---> Type extension);
+ val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0);
+ val truncate_decl = (truncateN, rec_schemeT0 --> recT0);
+
+ (* prepare definitions *)
+
+ fun parent_more s =
+ if null parents then s
+ else mk_sel s (Long_Name.qualify (#name (List.last parents)) moreN, extT);
+
+ fun parent_more_upd v s =
+ if null parents then v$s
+ else let val mp = Long_Name.qualify (#name (List.last parents)) moreN;
+ in mk_upd updateN mp v s end;
+
+ (*record (scheme) type abbreviation*)
+ val recordT_specs =
+ [(Binding.name (suffix schemeN bname), alphas @ [zeta], rec_schemeT0, Syntax.NoSyn),
+ (Binding.name bname, alphas, recT0, Syntax.NoSyn)];
+
+ (*selectors*)
+ fun mk_sel_spec (c,T) =
+ Const (mk_selC rec_schemeT0 (c,T))
+ :== (lambda r0 (Const (mk_selC extT (suffix ext_dest c,T))$parent_more r0));
+ val sel_specs = map mk_sel_spec fields_more;
+
+ (*updates*)
+
+ fun mk_upd_spec (c,T) =
+ let
+ val new = mk_upd' updN c (Free (base c,T-->T)) extT(*(parent_more r0)*);
+ in Const (mk_updC updateN rec_schemeT0 (c,T))$(Free (base c,T-->T))$r0
+ :== (parent_more_upd new r0)
+ end;
+ val upd_specs = map mk_upd_spec fields_more;
+
+ (*derived operations*)
+ val make_spec = Const (full makeN, all_types ---> recT0) $$ all_vars :==
+ mk_rec (all_vars @ [HOLogic.unit]) 0;
+ val fields_spec = Const (full fields_selN, types ---> Type extension) $$ vars :==
+ mk_rec (all_vars @ [HOLogic.unit]) parent_len;
+ val extend_spec =
+ Const (full extendN, recT0-->moreT-->rec_schemeT0) $ r_unit0 $ more :==
+ mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0;
+ val truncate_spec = Const (full truncateN, rec_schemeT0 --> recT0) $ r0 :==
+ mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0;
+
+ (* 2st stage: defs_thy *)
+
+ fun mk_defs () =
+ extension_thy
+ |> Sign.add_trfuns
+ ([],[],field_tr's, [])
+ |> Sign.add_advanced_trfuns
+ ([],[],adv_ext_tr's @ adv_record_type_tr's @ adv_record_type_abbr_tr's,[])
+ |> Sign.parent_path
+ |> Sign.add_tyabbrs_i recordT_specs
+ |> Sign.add_path bname
+ |> Sign.add_consts_i
+ (map2 (fn (x, T) => fn mx => (Binding.name x, T, mx))
+ sel_decls (field_syntax @ [Syntax.NoSyn]))
+ |> (Sign.add_consts_i o map (fn (x, T) => (Binding.name x, T, Syntax.NoSyn)))
+ (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
+ |> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) sel_specs)
+ ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) upd_specs)
+ ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name))
+ [make_spec, fields_spec, extend_spec, truncate_spec])
+ |-> (fn defs as ((sel_defs, upd_defs), derived_defs) =>
+ fold Code.add_default_eqn sel_defs
+ #> fold Code.add_default_eqn upd_defs
+ #> fold Code.add_default_eqn derived_defs
+ #> pair defs)
+ val (((sel_defs, upd_defs), derived_defs), defs_thy) =
+ timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:"
+ mk_defs;
+
+
+ (* prepare propositions *)
+ val _ = timing_msg "record preparing propositions";
+ val P = Free (Name.variant all_variants "P", rec_schemeT0-->HOLogic.boolT);
+ val C = Free (Name.variant all_variants "C", HOLogic.boolT);
+ val P_unit = Free (Name.variant all_variants "P", recT0-->HOLogic.boolT);
+
+ (*selectors*)
+ val sel_conv_props =
+ map (fn (c, x as Free (_,T)) => mk_sel r_rec0 (c,T) === x) named_vars_more;
+
+ (*updates*)
+ fun mk_upd_prop (i,(c,T)) =
+ let val x' = Free (Name.variant all_variants (base c ^ "'"),T-->T);
+ val n = parent_fields_len + i;
+ val args' = nth_map n (K (x'$nth all_vars_more n)) all_vars_more
+ in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end;
+ val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
+
+ (*induct*)
+ val induct_scheme_prop =
+ All (map dest_Free all_vars_more) (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0);
+ val induct_prop =
+ (All (map dest_Free all_vars) (Trueprop (P_unit $ r_rec_unit0)),
+ Trueprop (P_unit $ r_unit0));
+
+ (*surjective*)
+ val surjective_prop =
+ let val args = map (fn (c,Free (_,T)) => mk_sel r0 (c,T)) all_named_vars_more
+ in r0 === mk_rec args 0 end;
+
+ (*cases*)
+ val cases_scheme_prop =
+ (All (map dest_Free all_vars_more)
+ (Trueprop (HOLogic.mk_eq (r0,r_rec0)) ==> Trueprop C))
+ ==> Trueprop C;
+
+ val cases_prop =
+ (All (map dest_Free all_vars)
+ (Trueprop (HOLogic.mk_eq (r_unit0,r_rec_unit0)) ==> Trueprop C))
+ ==> Trueprop C;
+
+ (*split*)
+ val split_meta_prop =
+ let val P = Free (Name.variant all_variants "P", rec_schemeT0-->Term.propT) in
+ Logic.mk_equals
+ (All [dest_Free r0] (P $ r0), All (map dest_Free all_vars_more) (P $ r_rec0))
+ end;
+
+ val split_object_prop =
+ let fun ALL vs t = List.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) t vs
+ in (ALL [dest_Free r0] (P $ r0)) === (ALL (map dest_Free all_vars_more) (P $ r_rec0))
+ end;
+
+
+ val split_ex_prop =
+ let fun EX vs t = List.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) t vs
+ in (EX [dest_Free r0] (P $ r0)) === (EX (map dest_Free all_vars_more) (P $ r_rec0))
+ end;
+
+ (*equality*)
+ val equality_prop =
+ let
+ val s' = Free (rN ^ "'", rec_schemeT0)
+ fun mk_sel_eq (c,Free (_,T)) = mk_sel r0 (c,T) === mk_sel s' (c,T)
+ val seleqs = map mk_sel_eq all_named_vars_more
+ in All (map dest_Free [r0,s']) (Logic.list_implies (seleqs,r0 === s')) end;
+
+ (* 3rd stage: thms_thy *)
+
+ fun prove stndrd = quick_and_dirty_prove stndrd defs_thy;
+ val prove_standard = quick_and_dirty_prove true defs_thy;
+
+ fun prove_simp stndrd ss simps =
+ let val tac = simp_all_tac ss simps
+ in fn prop => prove stndrd [] prop (K tac) end;
+
+ val ss = get_simpset defs_thy;
+
+ fun sel_convs_prf () = map (prove_simp false ss
+ (sel_defs@ext_dest_convs)) sel_conv_props;
+ val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf;
+ fun sel_convs_standard_prf () = map standard sel_convs
+ val sel_convs_standard =
+ timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf;
+
+ fun upd_convs_prf () =
+ map (prove_simp true ss (upd_defs@u_convs)) upd_conv_props;
+
+ val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf;
+
+ val parent_induct = if null parents then [] else [#induct (hd (rev parents))];
+
+ fun induct_scheme_prf () = prove_standard [] induct_scheme_prop (fn _ =>
+ (EVERY [if null parent_induct
+ then all_tac else try_param_tac rN (hd parent_induct) 1,
+ try_param_tac rN ext_induct 1,
+ asm_simp_tac HOL_basic_ss 1]));
+ val induct_scheme = timeit_msg "record induct_scheme proof:" induct_scheme_prf;
+
+ fun induct_prf () =
+ let val (assm, concl) = induct_prop;
+ in
+ prove_standard [assm] concl (fn {prems, ...} =>
+ try_param_tac rN induct_scheme 1
+ THEN try_param_tac "more" @{thm unit.induct} 1
+ THEN resolve_tac prems 1)
+ end;
+ val induct = timeit_msg "record induct proof:" induct_prf;
+
+ fun surjective_prf () =
+ prove_standard [] surjective_prop (fn prems =>
+ (EVERY [try_param_tac rN induct_scheme 1,
+ simp_tac (ss addsimps sel_convs_standard) 1]))
+ val surjective = timeit_msg "record surjective proof:" surjective_prf;
+
+ fun cases_scheme_prf_opt () =
+ let
+ val (_$(Pvar$_)) = concl_of induct_scheme;
+ val ind = cterm_instantiate
+ [(cterm_of defs_thy Pvar, cterm_of defs_thy
+ (lambda w (HOLogic.imp$HOLogic.mk_eq(r0,w)$C)))]
+ induct_scheme;
+ in standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
+
+ fun cases_scheme_prf_noopt () =
+ prove_standard [] cases_scheme_prop (fn _ =>
+ EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1,
+ try_param_tac rN induct_scheme 1,
+ rtac impI 1,
+ REPEAT (etac allE 1),
+ etac mp 1,
+ rtac refl 1])
+ val cases_scheme_prf = quick_and_dirty_prf cases_scheme_prf_noopt cases_scheme_prf_opt;
+ val cases_scheme = timeit_msg "record cases_scheme proof:" cases_scheme_prf;
+
+ fun cases_prf () =
+ prove_standard [] cases_prop (fn _ =>
+ try_param_tac rN cases_scheme 1
+ THEN simp_all_tac HOL_basic_ss [unit_all_eq1]);
+ val cases = timeit_msg "record cases proof:" cases_prf;
+
+ fun split_meta_prf () =
+ prove false [] split_meta_prop (fn _ =>
+ EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
+ etac meta_allE 1, atac 1,
+ rtac (prop_subst OF [surjective]) 1,
+ REPEAT (etac meta_allE 1), atac 1]);
+ val split_meta = timeit_msg "record split_meta proof:" split_meta_prf;
+ val split_meta_standard = standard split_meta;
+
+ fun split_object_prf_opt () =
+ let
+ val cPI= cterm_of defs_thy (lambda r0 (Trueprop (P$r0)));
+ val (_$Abs(_,_,P$_)) = fst (Logic.dest_equals (concl_of split_meta_standard));
+ val cP = cterm_of defs_thy P;
+ val split_meta' = cterm_instantiate [(cP,cPI)] split_meta_standard;
+ val (l,r) = HOLogic.dest_eq (HOLogic.dest_Trueprop split_object_prop);
+ val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l);
+ val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r);
+ val thl = assume cl (*All r. P r*) (* 1 *)
+ |> obj_to_meta_all (*!!r. P r*)
+ |> equal_elim split_meta' (*!!n m more. P (ext n m more)*)
+ |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*)
+ |> implies_intr cl (* 1 ==> 2 *)
+ val thr = assume cr (*All n m more. P (ext n m more)*)
+ |> obj_to_meta_all (*!!n m more. P (ext n m more)*)
+ |> equal_elim (symmetric split_meta') (*!!r. P r*)
+ |> meta_to_obj_all (*All r. P r*)
+ |> implies_intr cr (* 2 ==> 1 *)
+ in standard (thr COMP (thl COMP iffI)) end;
+
+ fun split_object_prf_noopt () =
+ prove_standard [] split_object_prop (fn _ =>
+ EVERY [rtac iffI 1,
+ REPEAT (rtac allI 1), etac allE 1, atac 1,
+ rtac allI 1, rtac induct_scheme 1,REPEAT (etac allE 1),atac 1]);
+
+ val split_object_prf = quick_and_dirty_prf split_object_prf_noopt split_object_prf_opt;
+ val split_object = timeit_msg "record split_object proof:" split_object_prf;
+
+
+ fun split_ex_prf () =
+ prove_standard [] split_ex_prop (fn _ =>
+ EVERY [rtac iffI 1,
+ etac exE 1,
+ simp_tac (HOL_basic_ss addsimps [split_meta_standard]) 1,
+ ex_inst_tac 1,
+ (*REPEAT (rtac exI 1),*)
+ atac 1,
+ REPEAT (etac exE 1),
+ rtac exI 1,
+ atac 1]);
+ val split_ex = timeit_msg "record split_ex proof:" split_ex_prf;
+
+ fun equality_tac thms =
+ let val (s'::s::eqs) = rev thms;
+ val ss' = ss addsimps (s'::s::sel_convs_standard);
+ val eqs' = map (simplify ss') eqs;
+ in simp_tac (HOL_basic_ss addsimps (s'::s::eqs')) 1 end;
+
+ fun equality_prf () = prove_standard [] equality_prop (fn {context, ...} =>
+ fn st => let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in
+ st |> (res_inst_tac context [((rN, 0), s)] cases_scheme 1
+ THEN res_inst_tac context [((rN, 0), s')] cases_scheme 1
+ THEN (METAHYPS equality_tac 1))
+ (* simp_all_tac ss (sel_convs) would also work but is less efficient *)
+ end);
+ val equality = timeit_msg "record equality proof:" equality_prf;
+
+ val ((([sel_convs',upd_convs',sel_defs',upd_defs',[split_meta',split_object',split_ex'],derived_defs'],
+ [surjective',equality']),[induct_scheme',induct',cases_scheme',cases']), thms_thy) =
+ defs_thy
+ |> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
+ [("select_convs", sel_convs_standard),
+ ("update_convs", upd_convs),
+ ("select_defs", sel_defs),
+ ("update_defs", upd_defs),
+ ("splits", [split_meta_standard,split_object,split_ex]),
+ ("defs", derived_defs)]
+ ||>> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
+ [("surjective", surjective),
+ ("equality", equality)]
+ ||>> (PureThy.add_thms o (map o apfst o apfst) Binding.name)
+ [(("induct_scheme", induct_scheme), induct_type_global (suffix schemeN name)),
+ (("induct", induct), induct_type_global name),
+ (("cases_scheme", cases_scheme), cases_type_global (suffix schemeN name)),
+ (("cases", cases), cases_type_global name)];
+
+
+ val sel_upd_simps = sel_convs' @ upd_convs';
+ val iffs = [ext_inject]
+ val final_thy =
+ thms_thy
+ |> (snd oo PureThy.add_thmss)
+ [((Binding.name "simps", sel_upd_simps),
+ [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]),
+ ((Binding.name "iffs", iffs), [iff_add])]
+ |> put_record name (make_record_info args parent fields extension induct_scheme')
+ |> put_sel_upd (names @ [full_moreN]) sel_upd_simps
+ |> add_record_equalities extension_id equality'
+ |> add_extinjects ext_inject
+ |> add_extsplit extension_name ext_split
+ |> add_record_splits extension_id (split_meta',split_object',split_ex',induct_scheme')
+ |> add_extfields extension_name (fields @ [(full_moreN,moreT)])
+ |> add_fieldext (extension_name,snd extension) (names @ [full_moreN])
+ |> Sign.parent_path;
+
+ in final_thy
+ end;
+
+
+(* add_record *)
+
+(*we do all preparations and error checks here, deferring the real
+ work to record_definition*)
+fun gen_add_record prep_typ prep_raw_parent quiet_mode (params, bname) raw_parent raw_fields thy =
+ let
+ val _ = Theory.requires thy "Record" "record definitions";
+ val _ = if quiet_mode then () else writeln ("Defining record " ^ quote bname ^ " ...");
+
+ val ctxt = ProofContext.init thy;
+
+
+ (* parents *)
+
+ fun prep_inst T = fst (cert_typ ctxt T []);
+
+ val parent = Option.map (apfst (map prep_inst) o prep_raw_parent ctxt) raw_parent
+ handle ERROR msg => cat_error msg ("The error(s) above in parent record specification");
+ val parents = add_parents thy parent [];
+
+ val init_env =
+ (case parent of
+ NONE => []
+ | SOME (types, _) => List.foldr OldTerm.add_typ_tfrees [] types);
+
+
+ (* fields *)
+
+ fun prep_field (c, raw_T, mx) env =
+ let val (T, env') = prep_typ ctxt raw_T env handle ERROR msg =>
+ cat_error msg ("The error(s) above occured in record field " ^ quote c)
+ in ((c, T, mx), env') end;
+
+ val (bfields, envir) = fold_map prep_field raw_fields init_env;
+ val envir_names = map fst envir;
+
+
+ (* args *)
+
+ val defaultS = Sign.defaultS thy;
+ val args = map (fn x => (x, AList.lookup (op =) envir x |> the_default defaultS)) params;
+
+
+ (* errors *)
+
+ val name = Sign.full_bname thy bname;
+ val err_dup_record =
+ if is_none (get_record thy name) then []
+ else ["Duplicate definition of record " ^ quote name];
+
+ val err_dup_parms =
+ (case duplicates (op =) params of
+ [] => []
+ | dups => ["Duplicate parameter(s) " ^ commas dups]);
+
+ val err_extra_frees =
+ (case subtract (op =) params envir_names of
+ [] => []
+ | extras => ["Extra free type variable(s) " ^ commas extras]);
+
+ val err_no_fields = if null bfields then ["No fields present"] else [];
+
+ val err_dup_fields =
+ (case duplicates (op =) (map #1 bfields) of
+ [] => []
+ | dups => ["Duplicate field(s) " ^ commas_quote dups]);
+
+ val err_bad_fields =
+ if forall (not_equal moreN o #1) bfields then []
+ else ["Illegal field name " ^ quote moreN];
+
+ val err_dup_sorts =
+ (case duplicates (op =) envir_names of
+ [] => []
+ | dups => ["Inconsistent sort constraints for " ^ commas dups]);
+
+ val errs =
+ err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @
+ err_dup_fields @ err_bad_fields @ err_dup_sorts;
+ in
+ if null errs then () else error (cat_lines errs) ;
+ thy |> record_definition (args, bname) parent parents bfields
+ end
+ handle ERROR msg => cat_error msg ("Failed to define record " ^ quote bname);
+
+val add_record = gen_add_record read_typ read_raw_parent;
+val add_record_i = gen_add_record cert_typ (K I);
+
+(* setup theory *)
+
+val setup =
+ Sign.add_trfuns ([], parse_translation, [], []) #>
+ Sign.add_advanced_trfuns ([], adv_parse_translation, [], []) #>
+ Simplifier.map_simpset (fn ss =>
+ ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]);
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val record_decl =
+ P.type_args -- P.name --
+ (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const);
+
+val _ =
+ OuterSyntax.command "record" "define extensible record" K.thy_decl
+ (record_decl >> (fn (x, (y, z)) => Toplevel.theory (add_record false x y z)));
+
+end;
+
+end;
+
+
+structure BasicRecord: BASIC_RECORD = Record;
+open BasicRecord;
--- a/src/HOL/Tools/record_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2325 +0,0 @@
-(* Title: HOL/Tools/record_package.ML
- Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
-
-Extensible records with structural subtyping in HOL.
-*)
-
-
-signature BASIC_RECORD_PACKAGE =
-sig
- val record_simproc: simproc
- val record_eq_simproc: simproc
- val record_upd_simproc: simproc
- val record_split_simproc: (term -> int) -> simproc
- val record_ex_sel_eq_simproc: simproc
- val record_split_tac: int -> tactic
- val record_split_simp_tac: thm list -> (term -> int) -> int -> tactic
- val record_split_name: string
- val record_split_wrapper: string * wrapper
- val print_record_type_abbr: bool ref
- val print_record_type_as_fields: bool ref
-end;
-
-signature RECORD_PACKAGE =
-sig
- include BASIC_RECORD_PACKAGE
- val timing: bool ref
- val record_quick_and_dirty_sensitive: bool ref
- val updateN: string
- val updN: string
- val ext_typeN: string
- val extN: string
- val makeN: string
- val moreN: string
- val ext_dest: string
-
- val last_extT: typ -> (string * typ list) option
- val dest_recTs : typ -> (string * typ list) list
- val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ)
- val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ)
- val get_parent: theory -> string -> (typ list * string) option
- val get_extension: theory -> string -> (string * typ list) option
- val get_extinjects: theory -> thm list
- val get_simpset: theory -> simpset
- val print_records: theory -> unit
- val read_typ: Proof.context -> string -> (string * sort) list -> typ * (string * sort) list
- val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list
- val add_record: bool -> string list * string -> string option -> (string * string * mixfix) list
- -> theory -> theory
- val add_record_i: bool -> string list * string -> (typ list * string) option
- -> (string * typ * mixfix) list -> theory -> theory
- val setup: theory -> theory
-end;
-
-
-structure RecordPackage: RECORD_PACKAGE =
-struct
-
-val eq_reflection = thm "eq_reflection";
-val rec_UNIV_I = thm "rec_UNIV_I";
-val rec_True_simp = thm "rec_True_simp";
-val Pair_eq = thm "Product_Type.prod.inject";
-val atomize_all = thm "HOL.atomize_all";
-val atomize_imp = thm "HOL.atomize_imp";
-val meta_allE = thm "Pure.meta_allE";
-val prop_subst = thm "prop_subst";
-val Pair_sel_convs = [fst_conv,snd_conv];
-val K_record_comp = @{thm "K_record_comp"};
-val K_comp_convs = [@{thm o_apply}, K_record_comp]
-
-(** name components **)
-
-val rN = "r";
-val wN = "w";
-val moreN = "more";
-val schemeN = "_scheme";
-val ext_typeN = "_ext_type";
-val extN ="_ext";
-val casesN = "_cases";
-val ext_dest = "_sel";
-val updateN = "_update";
-val updN = "_upd";
-val makeN = "make";
-val fields_selN = "fields";
-val extendN = "extend";
-val truncateN = "truncate";
-
-(*see typedef_package.ML*)
-val RepN = "Rep_";
-val AbsN = "Abs_";
-
-(*** utilities ***)
-
-fun but_last xs = fst (split_last xs);
-
-fun varifyT midx =
- let fun varify (a, S) = TVar ((a, midx + 1), S);
- in map_type_tfree varify end;
-
-fun domain_type' T =
- domain_type T handle Match => T;
-
-fun range_type' T =
- range_type T handle Match => T;
-
-(* messages *)
-
-fun trace_thm str thm =
- tracing (str ^ (Pretty.string_of (Display.pretty_thm thm)));
-
-fun trace_thms str thms =
- (tracing str; map (trace_thm "") thms);
-
-fun trace_term str t =
- tracing (str ^ Syntax.string_of_term_global Pure.thy t);
-
-(* timing *)
-
-val timing = ref false;
-fun timeit_msg s x = if !timing then (warning s; timeit x) else x ();
-fun timing_msg s = if !timing then warning s else ();
-
-(* syntax *)
-
-fun prune n xs = Library.drop (n, xs);
-fun prefix_base s = Long_Name.map_base_name (fn bname => s ^ bname);
-
-val Trueprop = HOLogic.mk_Trueprop;
-fun All xs t = Term.list_all_free (xs, t);
-
-infix 9 $$;
-infix 0 :== ===;
-infixr 0 ==>;
-
-val (op $$) = Term.list_comb;
-val (op :==) = PrimitiveDefs.mk_defpair;
-val (op ===) = Trueprop o HOLogic.mk_eq;
-val (op ==>) = Logic.mk_implies;
-
-(* morphisms *)
-
-fun mk_RepN name = suffix ext_typeN (prefix_base RepN name);
-fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name);
-
-fun mk_Rep name repT absT =
- Const (suffix ext_typeN (prefix_base RepN name),absT --> repT);
-
-fun mk_Abs name repT absT =
- Const (mk_AbsN name,repT --> absT);
-
-(* constructor *)
-
-fun mk_extC (name,T) Ts = (suffix extN name, Ts ---> T);
-
-fun mk_ext (name,T) ts =
- let val Ts = map fastype_of ts
- in list_comb (Const (mk_extC (name,T) Ts),ts) end;
-
-(* cases *)
-
-fun mk_casesC (name,T,vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT)
-
-fun mk_cases (name,T,vT) f =
- let val Ts = binder_types (fastype_of f)
- in Const (mk_casesC (name,T,vT) Ts) $ f end;
-
-(* selector *)
-
-fun mk_selC sT (c,T) = (c,sT --> T);
-
-fun mk_sel s (c,T) =
- let val sT = fastype_of s
- in Const (mk_selC sT (c,T)) $ s end;
-
-(* updates *)
-
-fun mk_updC sfx sT (c,T) = (suffix sfx c, (T --> T) --> sT --> sT);
-
-fun mk_upd' sfx c v sT =
- let val vT = domain_type (fastype_of v);
- in Const (mk_updC sfx sT (c, vT)) $ v end;
-
-fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s
-
-(* types *)
-
-fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) =
- (case try (unsuffix ext_typeN) c_ext_type of
- NONE => raise TYPE ("RecordPackage.dest_recT", [typ], [])
- | SOME c => ((c, Ts), List.last Ts))
- | dest_recT typ = raise TYPE ("RecordPackage.dest_recT", [typ], []);
-
-fun is_recT T =
- (case try dest_recT T of NONE => false | SOME _ => true);
-
-fun dest_recTs T =
- let val ((c, Ts), U) = dest_recT T
- in (c, Ts) :: dest_recTs U
- end handle TYPE _ => [];
-
-fun last_extT T =
- let val ((c, Ts), U) = dest_recT T
- in (case last_extT U of
- NONE => SOME (c,Ts)
- | SOME l => SOME l)
- end handle TYPE _ => NONE
-
-fun rec_id i T =
- let val rTs = dest_recTs T
- val rTs' = if i < 0 then rTs else Library.take (i,rTs)
- in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end;
-
-(*** extend theory by record definition ***)
-
-(** record info **)
-
-(* type record_info and parent_info *)
-
-type record_info =
- {args: (string * sort) list,
- parent: (typ list * string) option,
- fields: (string * typ) list,
- extension: (string * typ list),
- induct: thm
- };
-
-fun make_record_info args parent fields extension induct =
- {args = args, parent = parent, fields = fields, extension = extension,
- induct = induct}: record_info;
-
-
-type parent_info =
- {name: string,
- fields: (string * typ) list,
- extension: (string * typ list),
- induct: thm
-};
-
-fun make_parent_info name fields extension induct =
- {name = name, fields = fields, extension = extension, induct = induct}: parent_info;
-
-
-(* theory data *)
-
-type record_data =
- {records: record_info Symtab.table,
- sel_upd:
- {selectors: unit Symtab.table,
- updates: string Symtab.table,
- simpset: Simplifier.simpset},
- equalities: thm Symtab.table,
- extinjects: thm list,
- extsplit: thm Symtab.table, (* maps extension name to split rule *)
- splits: (thm*thm*thm*thm) Symtab.table, (* !!,!,EX - split-equalities,induct rule *)
- extfields: (string*typ) list Symtab.table, (* maps extension to its fields *)
- fieldext: (string*typ list) Symtab.table (* maps field to its extension *)
-};
-
-fun make_record_data
- records sel_upd equalities extinjects extsplit splits extfields fieldext =
- {records = records, sel_upd = sel_upd,
- equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
- extfields = extfields, fieldext = fieldext }: record_data;
-
-structure RecordsData = TheoryDataFun
-(
- type T = record_data;
- val empty =
- make_record_data Symtab.empty
- {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss}
- Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty;
-
- val copy = I;
- val extend = I;
- fun merge _
- ({records = recs1,
- sel_upd = {selectors = sels1, updates = upds1, simpset = ss1},
- equalities = equalities1,
- extinjects = extinjects1,
- extsplit = extsplit1,
- splits = splits1,
- extfields = extfields1,
- fieldext = fieldext1},
- {records = recs2,
- sel_upd = {selectors = sels2, updates = upds2, simpset = ss2},
- equalities = equalities2,
- extinjects = extinjects2,
- extsplit = extsplit2,
- splits = splits2,
- extfields = extfields2,
- fieldext = fieldext2}) =
- make_record_data
- (Symtab.merge (K true) (recs1, recs2))
- {selectors = Symtab.merge (K true) (sels1, sels2),
- updates = Symtab.merge (K true) (upds1, upds2),
- simpset = Simplifier.merge_ss (ss1, ss2)}
- (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2))
- (Library.merge Thm.eq_thm_prop (extinjects1, extinjects2))
- (Symtab.merge Thm.eq_thm_prop (extsplit1,extsplit2))
- (Symtab.merge (fn ((a,b,c,d),(w,x,y,z))
- => Thm.eq_thm (a,w) andalso Thm.eq_thm (b,x) andalso
- Thm.eq_thm (c,y) andalso Thm.eq_thm (d,z))
- (splits1, splits2))
- (Symtab.merge (K true) (extfields1,extfields2))
- (Symtab.merge (K true) (fieldext1,fieldext2));
-);
-
-fun print_records thy =
- let
- val {records = recs, ...} = RecordsData.get thy;
- val prt_typ = Syntax.pretty_typ_global thy;
-
- fun pretty_parent NONE = []
- | pretty_parent (SOME (Ts, name)) =
- [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
-
- fun pretty_field (c, T) = Pretty.block
- [Pretty.str (Sign.extern_const thy c), Pretty.str " ::",
- Pretty.brk 1, Pretty.quote (prt_typ T)];
-
- fun pretty_record (name, {args, parent, fields, ...}: record_info) =
- Pretty.block (Pretty.fbreaks (Pretty.block
- [prt_typ (Type (name, map TFree args)), Pretty.str " = "] ::
- pretty_parent parent @ map pretty_field fields));
- in map pretty_record (Symtab.dest recs) |> Pretty.chunks |> Pretty.writeln end;
-
-
-(* access 'records' *)
-
-val get_record = Symtab.lookup o #records o RecordsData.get;
-
-fun put_record name info thy =
- let
- val {records, sel_upd, equalities, extinjects,extsplit,splits,extfields,fieldext} =
- RecordsData.get thy;
- val data = make_record_data (Symtab.update (name, info) records)
- sel_upd equalities extinjects extsplit splits extfields fieldext;
- in RecordsData.put data thy end;
-
-
-(* access 'sel_upd' *)
-
-val get_sel_upd = #sel_upd o RecordsData.get;
-
-val is_selector = Symtab.defined o #selectors o get_sel_upd;
-val get_updates = Symtab.lookup o #updates o get_sel_upd;
-fun get_simpset thy = Simplifier.theory_context thy (#simpset (get_sel_upd thy));
-
-fun put_sel_upd names simps = RecordsData.map (fn {records,
- sel_upd = {selectors, updates, simpset},
- equalities, extinjects, extsplit, splits, extfields, fieldext} =>
- make_record_data records
- {selectors = fold (fn name => Symtab.update (name, ())) names selectors,
- updates = fold (fn name => Symtab.update ((suffix updateN) name, name)) names updates,
- simpset = Simplifier.addsimps (simpset, simps)}
- equalities extinjects extsplit splits extfields fieldext);
-
-
-(* access 'equalities' *)
-
-fun add_record_equalities name thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
- RecordsData.get thy;
- val data = make_record_data records sel_upd
- (Symtab.update_new (name, thm) equalities) extinjects extsplit
- splits extfields fieldext;
- in RecordsData.put data thy end;
-
-val get_equalities =Symtab.lookup o #equalities o RecordsData.get;
-
-
-(* access 'extinjects' *)
-
-fun add_extinjects thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
- RecordsData.get thy;
- val data =
- make_record_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects) extsplit
- splits extfields fieldext;
- in RecordsData.put data thy end;
-
-val get_extinjects = rev o #extinjects o RecordsData.get;
-
-
-(* access 'extsplit' *)
-
-fun add_extsplit name thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
- RecordsData.get thy;
- val data = make_record_data records sel_upd
- equalities extinjects (Symtab.update_new (name, thm) extsplit) splits
- extfields fieldext;
- in RecordsData.put data thy end;
-
-val get_extsplit = Symtab.lookup o #extsplit o RecordsData.get;
-
-
-(* access 'splits' *)
-
-fun add_record_splits name thmP thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields,fieldext} =
- RecordsData.get thy;
- val data = make_record_data records sel_upd
- equalities extinjects extsplit (Symtab.update_new (name, thmP) splits)
- extfields fieldext;
- in RecordsData.put data thy end;
-
-val get_splits = Symtab.lookup o #splits o RecordsData.get;
-
-
-(* parent/extension of named record *)
-
-val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o RecordsData.get);
-val get_extension = Option.map #extension oo (Symtab.lookup o #records o RecordsData.get);
-
-
-(* access 'extfields' *)
-
-fun add_extfields name fields thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit,splits, extfields, fieldext} =
- RecordsData.get thy;
- val data = make_record_data records sel_upd
- equalities extinjects extsplit splits
- (Symtab.update_new (name, fields) extfields) fieldext;
- in RecordsData.put data thy end;
-
-val get_extfields = Symtab.lookup o #extfields o RecordsData.get;
-
-fun get_extT_fields thy T =
- let
- val ((name,Ts),moreT) = dest_recT T;
- val recname = let val (nm::recn::rst) = rev (Long_Name.explode name)
- in Long_Name.implode (rev (nm::rst)) end;
- val midx = maxidx_of_typs (moreT::Ts);
- val varifyT = varifyT midx;
- val {records,extfields,...} = RecordsData.get thy;
- val (flds,(more,_)) = split_last (Symtab.lookup_list extfields name);
- val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
-
- val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) (Vartab.empty);
- val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds;
- in (flds',(more,moreT)) end;
-
-fun get_recT_fields thy T =
- let
- val (root_flds,(root_more,root_moreT)) = get_extT_fields thy T;
- val (rest_flds,rest_more) =
- if is_recT root_moreT then get_recT_fields thy root_moreT
- else ([],(root_more,root_moreT));
- in (root_flds@rest_flds,rest_more) end;
-
-
-(* access 'fieldext' *)
-
-fun add_fieldext extname_types fields thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- RecordsData.get thy;
- val fieldext' =
- fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
- val data=make_record_data records sel_upd equalities extinjects extsplit
- splits extfields fieldext';
- in RecordsData.put data thy end;
-
-
-val get_fieldext = Symtab.lookup o #fieldext o RecordsData.get;
-
-
-(* parent records *)
-
-fun add_parents thy NONE parents = parents
- | add_parents thy (SOME (types, name)) parents =
- let
- fun err msg = error (msg ^ " parent record " ^ quote name);
-
- val {args, parent, fields, extension, induct} =
- (case get_record thy name of SOME info => info | NONE => err "Unknown");
- val _ = if length types <> length args then err "Bad number of arguments for" else ();
-
- fun bad_inst ((x, S), T) =
- if Sign.of_sort thy (T, S) then NONE else SOME x
- val bads = List.mapPartial bad_inst (args ~~ types);
- val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
-
- val inst = map fst args ~~ types;
- val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
- val parent' = Option.map (apfst (map subst)) parent;
- val fields' = map (apsnd subst) fields;
- val extension' = apsnd (map subst) extension;
- in
- add_parents thy parent'
- (make_parent_info name fields' extension' induct :: parents)
- end;
-
-
-
-(** concrete syntax for records **)
-
-(* decode type *)
-
-fun decode_type thy t =
- let
- fun get_sort xs n = AList.lookup (op =) xs (n: indexname) |> the_default (Sign.defaultS thy);
- val map_sort = Sign.intern_sort thy;
- in
- Syntax.typ_of_term (get_sort (Syntax.term_sorts map_sort t)) map_sort t
- |> Sign.intern_tycons thy
- end;
-
-
-(* parse translations *)
-
-fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) =
- if c = mark then Syntax.const (suffix sfx name) $ (Abs ("_",dummyT, arg))
- else raise TERM ("gen_field_tr: " ^ mark, [t])
- | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
-
-fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) =
- if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u
- else [gen_field_tr mark sfx tm]
- | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm];
-
-
-fun record_update_tr [t, u] =
- Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t)
- | record_update_tr ts = raise TERM ("record_update_tr", ts);
-
-fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts
- | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts
- | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
- (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts
- | update_name_tr ts = raise TERM ("update_name_tr", ts);
-
-fun dest_ext_field mark (t as (Const (c,_) $ Const (name,_) $ arg)) =
- if c = mark then (name,arg) else raise TERM ("dest_ext_field: " ^ mark, [t])
- | dest_ext_field _ t = raise TERM ("dest_ext_field", [t])
-
-fun dest_ext_fields sep mark (trm as (Const (c,_) $ t $ u)) =
- if c = sep then dest_ext_field mark t::dest_ext_fields sep mark u
- else [dest_ext_field mark trm]
- | dest_ext_fields _ mark t = [dest_ext_field mark t]
-
-fun gen_ext_fields_tr sep mark sfx more ctxt t =
- let
- val thy = ProofContext.theory_of ctxt;
- val msg = "error in record input: ";
- val fieldargs = dest_ext_fields sep mark t;
- fun splitargs (field::fields) ((name,arg)::fargs) =
- if can (unsuffix name) field
- then let val (args,rest) = splitargs fields fargs
- in (arg::args,rest) end
- else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
- | splitargs [] (fargs as (_::_)) = ([],fargs)
- | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t])
- | splitargs _ _ = ([],[]);
-
- fun mk_ext (fargs as (name,arg)::_) =
- (case get_fieldext thy (Sign.intern_const thy name) of
- SOME (ext,_) => (case get_extfields thy ext of
- SOME flds
- => let val (args,rest) =
- splitargs (map fst (but_last flds)) fargs;
- val more' = mk_ext rest;
- in list_comb (Syntax.const (suffix sfx ext),args@[more'])
- end
- | NONE => raise TERM(msg ^ "no fields defined for "
- ^ ext,[t]))
- | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
- | mk_ext [] = more
-
- in mk_ext fieldargs end;
-
-fun gen_ext_type_tr sep mark sfx more ctxt t =
- let
- val thy = ProofContext.theory_of ctxt;
- val msg = "error in record-type input: ";
- val fieldargs = dest_ext_fields sep mark t;
- fun splitargs (field::fields) ((name,arg)::fargs) =
- if can (unsuffix name) field
- then let val (args,rest) = splitargs fields fargs
- in (arg::args,rest) end
- else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
- | splitargs [] (fargs as (_::_)) = ([],fargs)
- | splitargs (_::_) [] = raise TERM (msg ^ "expecting more fields", [t])
- | splitargs _ _ = ([],[]);
-
- fun mk_ext (fargs as (name,arg)::_) =
- (case get_fieldext thy (Sign.intern_const thy name) of
- SOME (ext,alphas) =>
- (case get_extfields thy ext of
- SOME flds
- => (let
- val flds' = but_last flds;
- val types = map snd flds';
- val (args,rest) = splitargs (map fst flds') fargs;
- val argtypes = map (Sign.certify_typ thy o decode_type thy) args;
- val midx = fold (fn T => fn i => Int.max (maxidx_of_typ T, i))
- argtypes 0;
- val varifyT = varifyT midx;
- val vartypes = map varifyT types;
-
- val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes)
- Vartab.empty;
- val alphas' = map ((Syntax.term_of_typ (! Syntax.show_sorts)) o
- Envir.norm_type subst o varifyT)
- (but_last alphas);
-
- val more' = mk_ext rest;
- in list_comb (Syntax.const (suffix sfx ext),alphas'@[more'])
- end handle TYPE_MATCH => raise
- TERM (msg ^ "type is no proper record (extension)", [t]))
- | NONE => raise TERM (msg ^ "no fields defined for " ^ ext,[t]))
- | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
- | mk_ext [] = more
-
- in mk_ext fieldargs end;
-
-fun gen_adv_record_tr sep mark sfx unit ctxt [t] =
- gen_ext_fields_tr sep mark sfx unit ctxt t
- | gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
-
-fun gen_adv_record_scheme_tr sep mark sfx ctxt [t, more] =
- gen_ext_fields_tr sep mark sfx more ctxt t
- | gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
-
-fun gen_adv_record_type_tr sep mark sfx unit ctxt [t] =
- gen_ext_type_tr sep mark sfx unit ctxt t
- | gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
-
-fun gen_adv_record_type_scheme_tr sep mark sfx ctxt [t, more] =
- gen_ext_type_tr sep mark sfx more ctxt t
- | gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
-
-val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit;
-val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN;
-
-val adv_record_type_tr =
- gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN
- (Syntax.term_of_typ false (HOLogic.unitT));
-val adv_record_type_scheme_tr =
- gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN;
-
-
-val parse_translation =
- [("_record_update", record_update_tr),
- ("_update_name", update_name_tr)];
-
-
-val adv_parse_translation =
- [("_record",adv_record_tr),
- ("_record_scheme",adv_record_scheme_tr),
- ("_record_type",adv_record_type_tr),
- ("_record_type_scheme",adv_record_type_scheme_tr)];
-
-
-(* print translations *)
-
-val print_record_type_abbr = ref true;
-val print_record_type_as_fields = ref true;
-
-fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ k $ u) =
- let val t = (case k of (Abs (_,_,(Abs (_,_,t)$Bound 0)))
- => if null (loose_bnos t) then t else raise Match
- | Abs (x,_,t) => if null (loose_bnos t) then t else raise Match
- | _ => raise Match)
-
- (* (case k of (Const ("K_record",_)$t) => t
- | Abs (x,_,Const ("K_record",_)$t$Bound 0) => t
- | _ => raise Match)*)
- in
- (case try (unsuffix sfx) name_field of
- SOME name =>
- apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u)
- | NONE => ([], tm))
- end
- | gen_field_upds_tr' _ _ tm = ([], tm);
-
-fun record_update_tr' tm =
- let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in
- if null ts then raise Match
- else Syntax.const "_record_update" $ u $
- foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts)
- end;
-
-fun gen_field_tr' sfx tr' name =
- let val name_sfx = suffix sfx name
- in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end;
-
-fun record_tr' sep mark record record_scheme unit ctxt t =
- let
- val thy = ProofContext.theory_of ctxt;
- fun field_lst t =
- (case strip_comb t of
- (Const (ext,_),args as (_::_))
- => (case try (unsuffix extN) (Sign.intern_const thy ext) of
- SOME ext'
- => (case get_extfields thy ext' of
- SOME flds
- => (let
- val (f::fs) = but_last (map fst flds);
- val flds' = Sign.extern_const thy f :: map Long_Name.base_name fs;
- val (args',more) = split_last args;
- in (flds'~~args')@field_lst more end
- handle Library.UnequalLengths => [("",t)])
- | NONE => [("",t)])
- | NONE => [("",t)])
- | _ => [("",t)])
-
- val (flds,(_,more)) = split_last (field_lst t);
- val _ = if null flds then raise Match else ();
- val flds' = map (fn (n,t)=>Syntax.const mark$Syntax.const n$t) flds;
- val flds'' = foldr1 (fn (x,y) => Syntax.const sep$x$y) flds';
-
- in if unit more
- then Syntax.const record$flds''
- else Syntax.const record_scheme$flds''$more
- end
-
-fun gen_record_tr' name =
- let val name_sfx = suffix extN name;
- val unit = (fn Const (@{const_syntax "Product_Type.Unity"},_) => true | _ => false);
- fun tr' ctxt ts = record_tr' "_fields" "_field" "_record" "_record_scheme" unit ctxt
- (list_comb (Syntax.const name_sfx,ts))
- in (name_sfx,tr')
- end
-
-fun print_translation names =
- map (gen_field_tr' updateN record_update_tr') names;
-
-
-(* record_type_abbr_tr' tries to reconstruct the record name type abbreviation from *)
-(* the (nested) extension types. *)
-fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt tm =
- let
- val thy = ProofContext.theory_of ctxt;
- (* tm is term representation of a (nested) field type. We first reconstruct the *)
- (* type from tm so that we can continue on the type level rather then the term level.*)
-
- (* WORKAROUND:
- * If a record type occurs in an error message of type inference there
- * may be some internal frees donoted by ??:
- * (Const "_tfree",_)$Free ("??'a",_).
-
- * This will unfortunately be translated to Type ("??'a",[]) instead of
- * TFree ("??'a",_) by typ_of_term, which will confuse unify below.
- * fixT works around.
- *)
- fun fixT (T as Type (x,[])) =
- if String.isPrefix "??'" x then TFree (x,Sign.defaultS thy) else T
- | fixT (Type (x,xs)) = Type (x,map fixT xs)
- | fixT T = T;
-
- val T = fixT (decode_type thy tm);
- val midx = maxidx_of_typ T;
- val varifyT = varifyT midx;
-
- fun mk_type_abbr subst name alphas =
- let val abbrT = Type (name, map (fn a => varifyT (TFree (a, Sign.defaultS thy))) alphas);
- in Syntax.term_of_typ (! Syntax.show_sorts)
- (Sign.extern_typ thy (Envir.norm_type subst abbrT)) end;
-
- fun match rT T = (Sign.typ_match thy (varifyT rT,T)
- Vartab.empty);
-
- in if !print_record_type_abbr
- then (case last_extT T of
- SOME (name,_)
- => if name = lastExt
- then
- (let
- val subst = match schemeT T
- in
- if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree(zeta,Sign.defaultS thy))))
- then mk_type_abbr subst abbr alphas
- else mk_type_abbr subst (suffix schemeN abbr) (alphas@[zeta])
- end handle TYPE_MATCH => default_tr' ctxt tm)
- else raise Match (* give print translation of specialised record a chance *)
- | _ => raise Match)
- else default_tr' ctxt tm
- end
-
-fun record_type_tr' sep mark record record_scheme ctxt t =
- let
- val thy = ProofContext.theory_of ctxt;
-
- val T = decode_type thy t;
- val varifyT = varifyT (Term.maxidx_of_typ T);
-
- fun term_of_type T = Syntax.term_of_typ (!Syntax.show_sorts) (Sign.extern_typ thy T);
-
- fun field_lst T =
- (case T of
- Type (ext, args)
- => (case try (unsuffix ext_typeN) ext of
- SOME ext'
- => (case get_extfields thy ext' of
- SOME flds
- => (case get_fieldext thy (fst (hd flds)) of
- SOME (_, alphas)
- => (let
- val (f :: fs) = but_last flds;
- val flds' = apfst (Sign.extern_const thy) f
- :: map (apfst Long_Name.base_name) fs;
- val (args', more) = split_last args;
- val alphavars = map varifyT (but_last alphas);
- val subst = fold2 (curry (Sign.typ_match thy))
- alphavars args' Vartab.empty;
- val flds'' = (map o apsnd)
- (Envir.norm_type subst o varifyT) flds';
- in flds'' @ field_lst more end
- handle TYPE_MATCH => [("", T)]
- | Library.UnequalLengths => [("", T)])
- | NONE => [("", T)])
- | NONE => [("", T)])
- | NONE => [("", T)])
- | _ => [("", T)])
-
- val (flds, (_, moreT)) = split_last (field_lst T);
- val flds' = map (fn (n, T) => Syntax.const mark $ Syntax.const n $ term_of_type T) flds;
- val flds'' = foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds' handle Empty => raise Match;
-
- in if not (!print_record_type_as_fields) orelse null flds then raise Match
- else if moreT = HOLogic.unitT
- then Syntax.const record$flds''
- else Syntax.const record_scheme$flds''$term_of_type moreT
- end
-
-
-fun gen_record_type_tr' name =
- let val name_sfx = suffix ext_typeN name;
- fun tr' ctxt ts = record_type_tr' "_field_types" "_field_type"
- "_record_type" "_record_type_scheme" ctxt
- (list_comb (Syntax.const name_sfx,ts))
- in (name_sfx,tr')
- end
-
-
-fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name =
- let val name_sfx = suffix ext_typeN name;
- val default_tr' = record_type_tr' "_field_types" "_field_type"
- "_record_type" "_record_type_scheme"
- fun tr' ctxt ts =
- record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt
- (list_comb (Syntax.const name_sfx,ts))
- in (name_sfx, tr') end;
-
-(** record simprocs **)
-
-val record_quick_and_dirty_sensitive = ref false;
-
-
-fun quick_and_dirty_prove stndrd thy asms prop tac =
- if !record_quick_and_dirty_sensitive andalso !quick_and_dirty
- then Goal.prove (ProofContext.init thy) [] []
- (Logic.list_implies (map Logic.varify asms,Logic.varify prop))
- (K (SkipProof.cheat_tac @{theory HOL}))
- (* standard can take quite a while for large records, thats why
- * we varify the proposition manually here.*)
- else let val prf = Goal.prove (ProofContext.init thy) [] asms prop tac;
- in if stndrd then standard prf else prf end;
-
-fun quick_and_dirty_prf noopt opt () =
- if !record_quick_and_dirty_sensitive andalso !quick_and_dirty
- then noopt ()
- else opt ();
-
-local
-fun abstract_over_fun_app (Abs (f,fT,t)) =
- let
- val (f',t') = Term.dest_abs (f,fT,t);
- val T = domain_type fT;
- val (x,T') = hd (Term.variant_frees t' [("x",T)]);
- val f_x = Free (f',fT)$(Free (x,T'));
- fun is_constr (Const (c,_)$_) = can (unsuffix extN) c
- | is_constr _ = false;
- fun subst (t as u$w) = if Free (f',fT)=u
- then if is_constr w then f_x
- else raise TERM ("abstract_over_fun_app",[t])
- else subst u$subst w
- | subst (Abs (x,T,t)) = (Abs (x,T,subst t))
- | subst t = t
- val t'' = abstract_over (f_x,subst t');
- val vars = strip_qnt_vars "all" t'';
- val bdy = strip_qnt_body "all" t'';
-
- in list_abs ((x,T')::vars,bdy) end
- | abstract_over_fun_app t = raise TERM ("abstract_over_fun_app",[t]);
-(* Generates a theorem of the kind:
- * !!f x*. PROP P (f ( r x* ) x* == !!r x*. PROP P r x*
- *)
-fun mk_fun_apply_eq (Abs (f, fT, t)) thy =
- let
- val rT = domain_type fT;
- val vars = Term.strip_qnt_vars "all" t;
- val Ts = map snd vars;
- val n = length vars;
- fun app_bounds 0 t = t$Bound 0
- | app_bounds n t = if n > 0 then app_bounds (n-1) (t$Bound n) else t
-
-
- val [P,r] = Term.variant_frees t [("P",rT::Ts--->Term.propT),("r",Ts--->rT)];
- val prop = Logic.mk_equals
- (list_all ((f,fT)::vars,
- app_bounds (n - 1) ((Free P)$(Bound n$app_bounds (n-1) (Free r)))),
- list_all ((fst r,rT)::vars,
- app_bounds (n - 1) ((Free P)$Bound n)));
- val prove_standard = quick_and_dirty_prove true thy;
- val thm = prove_standard [] prop (fn _ =>
- EVERY [rtac equal_intr_rule 1,
- Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1,
- Goal.norm_hhf_tac 1,REPEAT (etac meta_allE 1), atac 1]);
- in thm end
- | mk_fun_apply_eq t thy = raise TERM ("mk_fun_apply_eq",[t]);
-
-in
-(* During proof of theorems produced by record_simproc you can end up in
- * situations like "!!f ... . ... f r ..." where f is an extension update function.
- * In order to split "f r" we transform this to "!!r ... . ... r ..." so that the
- * usual split rules for extensions can apply.
- *)
-val record_split_f_more_simproc =
- Simplifier.simproc @{theory HOL} "record_split_f_more_simp" ["x"]
- (fn thy => fn _ => fn t =>
- (case t of (Const ("all", Type (_, [Type (_, [Type("fun",[T,T']), _]), _])))$
- (trm as Abs _) =>
- (case rec_id (~1) T of
- "" => NONE
- | n => if T=T'
- then (let
- val P=cterm_of thy (abstract_over_fun_app trm);
- val thm = mk_fun_apply_eq trm thy;
- val PV = cterm_of thy (hd (OldTerm.term_vars (prop_of thm)));
- val thm' = cterm_instantiate [(PV,P)] thm;
- in SOME thm' end handle TERM _ => NONE)
- else NONE)
- | _ => NONE))
-end
-
-fun prove_split_simp thy ss T prop =
- let
- val {sel_upd={simpset,...},extsplit,...} = RecordsData.get thy;
- val extsplits =
- Library.foldl (fn (thms,(n,_)) => the_list (Symtab.lookup extsplit n) @ thms)
- ([],dest_recTs T);
- val thms = (case get_splits thy (rec_id (~1) T) of
- SOME (all_thm,_,_,_) =>
- all_thm::(case extsplits of [thm] => [] | _ => extsplits)
- (* [thm] is the same as all_thm *)
- | NONE => extsplits)
- val thms'=K_comp_convs@thms;
- val ss' = (Simplifier.inherit_context ss simpset
- addsimps thms'
- addsimprocs [record_split_f_more_simproc]);
- in
- quick_and_dirty_prove true thy [] prop (fn _ => simp_tac ss' 1)
- end;
-
-
-local
-fun eq (s1:string) (s2:string) = (s1 = s2);
-fun has_field extfields f T =
- exists (fn (eN,_) => exists (eq f o fst) (Symtab.lookup_list extfields eN))
- (dest_recTs T);
-
-fun K_skeleton n (T as Type (_,[_,kT])) (b as Bound i) (Abs (x,xT,t)) =
- if null (loose_bnos t) then ((n,kT),(Abs (x,xT,Bound (i+1)))) else ((n,T),b)
- | K_skeleton n T b _ = ((n,T),b);
-
-(*
-fun K_skeleton n _ b ((K_rec as Const ("Record.K_record",Type (_,[kT,_])))$_) =
- ((n,kT),K_rec$b)
- | K_skeleton n _ (Bound i)
- (Abs (x,T,(K_rec as Const ("Record.K_record",Type (_,[kT,_])))$_$Bound 0)) =
- ((n,kT),Abs (x,T,(K_rec$Bound (i+1)$Bound 0)))
- | K_skeleton n T b _ = ((n,T),b);
- *)
-
-fun normalize_rhs thm =
- let
- val ss = HOL_basic_ss addsimps K_comp_convs;
- val rhs = thm |> Thm.cprop_of |> Thm.dest_comb |> snd;
- val rhs' = (Simplifier.rewrite ss rhs);
- in Thm.transitive thm rhs' end;
-in
-(* record_simproc *)
-(* Simplifies selections of an record update:
- * (1) S (S_update k r) = k (S r)
- * (2) S (X_update k r) = S r
- * The simproc skips multiple updates at once, eg:
- * S (X_update x (Y_update y (S_update k r))) = k (S r)
- * But be careful in (2) because of the extendibility of records.
- * - If S is a more-selector we have to make sure that the update on component
- * X does not affect the selected subrecord.
- * - If X is a more-selector we have to make sure that S is not in the updated
- * subrecord.
- *)
-val record_simproc =
- Simplifier.simproc @{theory HOL} "record_simp" ["x"]
- (fn thy => fn ss => fn t =>
- (case t of (sel as Const (s, Type (_,[domS,rangeS])))$
- ((upd as Const (u,Type(_,[_,Type (_,[rT,_])]))) $ k $ r)=>
- if is_selector thy s then
- (case get_updates thy u of SOME u_name =>
- let
- val {sel_upd={updates,...},extfields,...} = RecordsData.get thy;
-
- fun mk_eq_terms ((upd as Const (u,Type(_,[kT,_]))) $ k $ r) =
- (case Symtab.lookup updates u of
- NONE => NONE
- | SOME u_name
- => if u_name = s
- then (case mk_eq_terms r of
- NONE =>
- let
- val rv = ("r",rT)
- val rb = Bound 0
- val (kv,kb) = K_skeleton "k" kT (Bound 1) k;
- in SOME (upd$kb$rb,kb$(sel$rb),[kv,rv]) end
- | SOME (trm,trm',vars) =>
- let
- val (kv,kb) = K_skeleton "k" kT (Bound (length vars)) k;
- in SOME (upd$kb$trm,kb$trm',kv::vars) end)
- else if has_field extfields u_name rangeS
- orelse has_field extfields s (domain_type kT)
- then NONE
- else (case mk_eq_terms r of
- SOME (trm,trm',vars)
- => let
- val (kv,kb) =
- K_skeleton "k" kT (Bound (length vars)) k;
- in SOME (upd$kb$trm,trm',kv::vars) end
- | NONE
- => let
- val rv = ("r",rT)
- val rb = Bound 0
- val (kv,kb) = K_skeleton "k" kT (Bound 1) k;
- in SOME (upd$kb$rb,sel$rb,[kv,rv]) end))
- | mk_eq_terms r = NONE
- in
- (case mk_eq_terms (upd$k$r) of
- SOME (trm,trm',vars)
- => SOME (prove_split_simp thy ss domS
- (list_all(vars, Logic.mk_equals (sel $ trm, trm'))))
- | NONE => NONE)
- end
- | NONE => NONE)
- else NONE
- | _ => NONE));
-
-(* record_upd_simproc *)
-(* simplify multiple updates:
- * (1) "N_update y (M_update g (N_update x (M_update f r))) =
- (N_update (y o x) (M_update (g o f) r))"
- * (2) "r(|M:= M r|) = r"
- * For (2) special care of "more" updates has to be taken:
- * r(|more := m; A := A r|)
- * If A is contained in the fields of m we cannot remove the update A := A r!
- * (But r(|more := r; A := A (r(|more := r|))|) = r(|more := r|)
-*)
-val record_upd_simproc =
- Simplifier.simproc @{theory HOL} "record_upd_simp" ["x"]
- (fn thy => fn ss => fn t =>
- (case t of ((upd as Const (u, Type(_,[_,Type(_,[rT,_])]))) $ k $ r) =>
- let datatype ('a,'b) calc = Init of 'b | Inter of 'a
- val {sel_upd={selectors,updates,...},extfields,...} = RecordsData.get thy;
-
- (*fun mk_abs_var x t = (x, fastype_of t);*)
- fun sel_name u = Long_Name.base_name (unsuffix updateN u);
-
- fun seed s (upd as Const (more,Type(_,[mT,_]))$ k $ r) =
- if has_field extfields s (domain_type' mT) then upd else seed s r
- | seed _ r = r;
-
- fun grow u uT k kT vars (sprout,skeleton) =
- if sel_name u = moreN
- then let val (kv,kb) = K_skeleton "k" kT (Bound (length vars)) k;
- in ((Const (u,uT)$k$sprout,Const (u,uT)$kb$skeleton),kv::vars) end
- else ((sprout,skeleton),vars);
-
-
- fun dest_k (Abs (x,T,((sel as Const (s,_))$r))) =
- if null (loose_bnos r) then SOME (x,T,sel,s,r) else NONE
- | dest_k (Abs (_,_,(Abs (x,T,((sel as Const (s,_))$r)))$Bound 0)) =
- (* eta expanded variant *)
- if null (loose_bnos r) then SOME (x,T,sel,s,r) else NONE
- | dest_k _ = NONE;
-
- fun is_upd_same (sprout,skeleton) u k =
- (case dest_k k of SOME (x,T,sel,s,r) =>
- if (unsuffix updateN u) = s andalso (seed s sprout) = r
- then SOME (fn t => Abs (x,T,incr_boundvars 1 t),sel,seed s skeleton)
- else NONE
- | NONE => NONE);
-
- fun init_seed r = ((r,Bound 0), [("r", rT)]);
-
- fun add (n:string) f fmaps =
- (case AList.lookup (op =) fmaps n of
- NONE => AList.update (op =) (n,[f]) fmaps
- | SOME fs => AList.update (op =) (n,f::fs) fmaps)
-
- fun comps (n:string) T fmaps =
- (case AList.lookup (op =) fmaps n of
- SOME fs =>
- foldr1 (fn (f,g) => Const ("Fun.comp",(T-->T)-->(T-->T)-->(T-->T))$f$g) fs
- | NONE => error ("record_upd_simproc.comps"))
-
- (* mk_updterm returns either
- * - Init (orig-term, orig-term-skeleton, vars) if no optimisation can be made,
- * where vars are the bound variables in the skeleton
- * - Inter (orig-term-skeleton,simplified-term-skeleton,
- * vars, (term-sprout, skeleton-sprout))
- * where "All vars. orig-term-skeleton = simplified-term-skeleton" is
- * the desired simplification rule,
- * the sprouts accumulate the "more-updates" on the way from the seed
- * to the outermost update. It is only relevant to calculate the
- * possible simplification for (2)
- * The algorithm first walks down the updates to the seed-record while
- * memorising the updates in the already-table. While walking up the
- * updates again, the optimised term is constructed.
- *)
- fun mk_updterm upds already
- (t as ((upd as Const (u,uT as (Type (_,[kT,_])))) $ k $ r)) =
- if Symtab.defined upds u
- then let
- fun rest already = mk_updterm upds already
- in if u mem_string already
- then (case (rest already r) of
- Init ((sprout,skel),vars) =>
- let
- val n = sel_name u;
- val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
- val (sprout',vars')= grow u uT k kT (kv::vars) (sprout,skel);
- in Inter (upd$kb$skel,skel,vars',add n kb [],sprout') end
- | Inter (trm,trm',vars,fmaps,sprout) =>
- let
- val n = sel_name u;
- val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
- val (sprout',vars') = grow u uT k kT (kv::vars) sprout;
- in Inter(upd$kb$trm,trm',kv::vars',add n kb fmaps,sprout')
- end)
- else
- (case rest (u::already) r of
- Init ((sprout,skel),vars) =>
- (case is_upd_same (sprout,skel) u k of
- SOME (K_rec,sel,skel') =>
- let
- val (sprout',vars') = grow u uT k kT vars (sprout,skel);
- in Inter(upd$(K_rec (sel$skel'))$skel,skel,vars',[],sprout')
- end
- | NONE =>
- let
- val n = sel_name u;
- val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
- in Init ((upd$k$sprout,upd$kb$skel),kv::vars) end)
- | Inter (trm,trm',vars,fmaps,sprout) =>
- (case is_upd_same sprout u k of
- SOME (K_rec,sel,skel) =>
- let
- val (sprout',vars') = grow u uT k kT vars sprout
- in Inter(upd$(K_rec (sel$skel))$trm,trm',vars',fmaps,sprout')
- end
- | NONE =>
- let
- val n = sel_name u
- val T = domain_type kT
- val (kv,kb) = K_skeleton n kT (Bound (length vars)) k;
- val (sprout',vars') = grow u uT k kT (kv::vars) sprout
- val fmaps' = add n kb fmaps
- in Inter (upd$kb$trm,upd$comps n T fmaps'$trm'
- ,vars',fmaps',sprout') end))
- end
- else Init (init_seed t)
- | mk_updterm _ _ t = Init (init_seed t);
-
- in (case mk_updterm updates [] t of
- Inter (trm,trm',vars,_,_)
- => SOME (normalize_rhs
- (prove_split_simp thy ss rT
- (list_all(vars, Logic.mk_equals (trm, trm')))))
- | _ => NONE)
- end
- | _ => NONE))
-end
-
-(* record_eq_simproc *)
-(* looks up the most specific record-equality.
- * Note on efficiency:
- * Testing equality of records boils down to the test of equality of all components.
- * Therefore the complexity is: #components * complexity for single component.
- * Especially if a record has a lot of components it may be better to split up
- * the record first and do simplification on that (record_split_simp_tac).
- * e.g. r(|lots of updates|) = x
- *
- * record_eq_simproc record_split_simp_tac
- * Complexity: #components * #updates #updates
- *
- *)
-val record_eq_simproc =
- Simplifier.simproc @{theory HOL} "record_eq_simp" ["r = s"]
- (fn thy => fn _ => fn t =>
- (case t of Const ("op =", Type (_, [T, _])) $ _ $ _ =>
- (case rec_id (~1) T of
- "" => NONE
- | name => (case get_equalities thy name of
- NONE => NONE
- | SOME thm => SOME (thm RS Eq_TrueI)))
- | _ => NONE));
-
-(* record_split_simproc *)
-(* splits quantified occurrences of records, for which P holds. P can peek on the
- * subterm starting at the quantified occurrence of the record (including the quantifier)
- * P t = 0: do not split
- * P t = ~1: completely split
- * P t > 0: split up to given bound of record extensions
- *)
-fun record_split_simproc P =
- Simplifier.simproc @{theory HOL} "record_split_simp" ["x"]
- (fn thy => fn _ => fn t =>
- (case t of (Const (quantifier, Type (_, [Type (_, [T, _]), _])))$trm =>
- if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex"
- then (case rec_id (~1) T of
- "" => NONE
- | name
- => let val split = P t
- in if split <> 0 then
- (case get_splits thy (rec_id split T) of
- NONE => NONE
- | SOME (all_thm, All_thm, Ex_thm,_)
- => SOME (case quantifier of
- "all" => all_thm
- | "All" => All_thm RS eq_reflection
- | "Ex" => Ex_thm RS eq_reflection
- | _ => error "record_split_simproc"))
- else NONE
- end)
- else NONE
- | _ => NONE))
-
-val record_ex_sel_eq_simproc =
- Simplifier.simproc @{theory HOL} "record_ex_sel_eq_simproc" ["Ex t"]
- (fn thy => fn ss => fn t =>
- let
- fun prove prop =
- quick_and_dirty_prove true thy [] prop
- (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy)
- addsimps simp_thms addsimprocs [record_split_simproc (K ~1)]) 1);
-
- fun mkeq (lr,Teq,(sel,Tsel),x) i =
- if is_selector thy sel then
- let val x' = if not (loose_bvar1 (x,0))
- then Free ("x" ^ string_of_int i, range_type Tsel)
- else raise TERM ("",[x]);
- val sel' = Const (sel,Tsel)$Bound 0;
- val (l,r) = if lr then (sel',x') else (x',sel');
- in Const ("op =",Teq)$l$r end
- else raise TERM ("",[Const (sel,Tsel)]);
-
- fun dest_sel_eq (Const ("op =",Teq)$(Const (sel,Tsel)$Bound 0)$X) =
- (true,Teq,(sel,Tsel),X)
- | dest_sel_eq (Const ("op =",Teq)$X$(Const (sel,Tsel)$Bound 0)) =
- (false,Teq,(sel,Tsel),X)
- | dest_sel_eq _ = raise TERM ("",[]);
-
- in
- (case t of
- (Const ("Ex",Tex)$Abs(s,T,t)) =>
- (let val eq = mkeq (dest_sel_eq t) 0;
- val prop = list_all ([("r",T)],
- Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
- HOLogic.true_const));
- in SOME (prove prop) end
- handle TERM _ => NONE)
- | _ => NONE)
- end)
-
-
-
-
-local
-val inductive_atomize = thms "induct_atomize";
-val inductive_rulify = thms "induct_rulify";
-in
-(* record_split_simp_tac *)
-(* splits (and simplifies) all records in the goal for which P holds.
- * For quantified occurrences of a record
- * P can peek on the whole subterm (including the quantifier); for free variables P
- * can only peek on the variable itself.
- * P t = 0: do not split
- * P t = ~1: completely split
- * P t > 0: split up to given bound of record extensions
- *)
-fun record_split_simp_tac thms P i st =
- let
- val thy = Thm.theory_of_thm st;
-
- val has_rec = exists_Const
- (fn (s, Type (_, [Type (_, [T, _]), _])) =>
- (s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T
- | _ => false);
-
- val goal = nth (Thm.prems_of st) (i - 1);
- val frees = List.filter (is_recT o type_of) (OldTerm.term_frees goal);
-
- fun mk_split_free_tac free induct_thm i =
- let val cfree = cterm_of thy free;
- val (_$(_$r)) = concl_of induct_thm;
- val crec = cterm_of thy r;
- val thm = cterm_instantiate [(crec,cfree)] induct_thm;
- in EVERY [simp_tac (HOL_basic_ss addsimps inductive_atomize) i,
- rtac thm i,
- simp_tac (HOL_basic_ss addsimps inductive_rulify) i]
- end;
-
- fun split_free_tac P i (free as Free (n,T)) =
- (case rec_id (~1) T of
- "" => NONE
- | name => let val split = P free
- in if split <> 0 then
- (case get_splits thy (rec_id split T) of
- NONE => NONE
- | SOME (_,_,_,induct_thm)
- => SOME (mk_split_free_tac free induct_thm i))
- else NONE
- end)
- | split_free_tac _ _ _ = NONE;
-
- val split_frees_tacs = List.mapPartial (split_free_tac P i) frees;
-
- val simprocs = if has_rec goal then [record_split_simproc P] else [];
- val thms' = K_comp_convs@thms
- in st |> ((EVERY split_frees_tacs)
- THEN (Simplifier.full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i))
- end handle Empty => Seq.empty;
-end;
-
-
-(* record_split_tac *)
-(* splits all records in the goal, which are quantified by ! or !!. *)
-fun record_split_tac i st =
- let
- val thy = Thm.theory_of_thm st;
-
- val has_rec = exists_Const
- (fn (s, Type (_, [Type (_, [T, _]), _])) =>
- (s = "all" orelse s = "All") andalso is_recT T
- | _ => false);
-
- val goal = nth (Thm.prems_of st) (i - 1);
-
- fun is_all t =
- (case t of (Const (quantifier, _)$_) =>
- if quantifier = "All" orelse quantifier = "all" then ~1 else 0
- | _ => 0);
-
- in if has_rec goal
- then Simplifier.full_simp_tac
- (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i st
- else Seq.empty
- end handle Subscript => Seq.empty;
-
-(* wrapper *)
-
-val record_split_name = "record_split_tac";
-val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac);
-
-
-
-(** theory extender interface **)
-
-(* prepare arguments *)
-
-fun read_raw_parent ctxt raw_T =
- (case ProofContext.read_typ_abbrev ctxt raw_T of
- Type (name, Ts) => (Ts, name)
- | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
-
-fun read_typ ctxt raw_T env =
- let
- val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;
- val T = Syntax.read_typ ctxt' raw_T;
- val env' = OldTerm.add_typ_tfrees (T, env);
- in (T, env') end;
-
-fun cert_typ ctxt raw_T env =
- let
- val thy = ProofContext.theory_of ctxt;
- val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg;
- val env' = OldTerm.add_typ_tfrees (T, env);
- in (T, env') end;
-
-
-(* attributes *)
-
-fun case_names_fields x = RuleCases.case_names ["fields"] x;
-fun induct_type_global name = [case_names_fields, Induct.induct_type name];
-fun cases_type_global name = [case_names_fields, Induct.cases_type name];
-
-(* tactics *)
-
-fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps));
-
-(* do case analysis / induction according to rule on last parameter of ith subgoal
- * (or on s if there are no parameters);
- * Instatiation of record variable (and predicate) in rule is calculated to
- * avoid problems with higher order unification.
- *)
-
-fun try_param_tac s rule i st =
- let
- val cert = cterm_of (Thm.theory_of_thm st);
- val g = nth (prems_of st) (i - 1);
- val params = Logic.strip_params g;
- val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g);
- val rule' = Thm.lift_rule (Thm.cprem_of st i) rule;
- val (P, ys) = strip_comb (HOLogic.dest_Trueprop
- (Logic.strip_assums_concl (prop_of rule')));
- (* ca indicates if rule is a case analysis or induction rule *)
- val (x, ca) = (case rev (Library.drop (length params, ys)) of
- [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
- (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
- | [x] => (head_of x, false));
- val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of
- [] => (case AList.lookup (op =) (map dest_Free (OldTerm.term_frees (prop_of st))) s of
- NONE => sys_error "try_param_tac: no such variable"
- | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl),
- (x, Free (s, T))])
- | (_, T) :: _ => [(P, list_abs (params, if ca then concl
- else incr_boundvars 1 (Abs (s, T, concl)))),
- (x, list_abs (params, Bound 0))])) rule'
- in compose_tac (false, rule'', nprems_of rule) i st end;
-
-
-(* !!x1 ... xn. ... ==> EX x1 ... xn. P x1 ... xn;
- instantiates x1 ... xn with parameters x1 ... xn *)
-fun ex_inst_tac i st =
- let
- val thy = Thm.theory_of_thm st;
- val g = nth (prems_of st) (i - 1);
- val params = Logic.strip_params g;
- val exI' = Thm.lift_rule (Thm.cprem_of st i) exI;
- val (_$(_$x)) = Logic.strip_assums_concl (hd (prems_of exI'));
- val cx = cterm_of thy (fst (strip_comb x));
-
- in Seq.single (Library.foldl (fn (st,v) =>
- Seq.hd
- (compose_tac (false, cterm_instantiate
- [(cx,cterm_of thy (list_abs (params,Bound v)))] exI',1)
- i st)) (st,((length params) - 1) downto 0))
- end;
-
-fun extension_typedef name repT alphas thy =
- let
- fun get_thms thy name =
- let
- val SOME { Abs_induct = abs_induct,
- Abs_inject=abs_inject, Abs_inverse = abs_inverse, ...} = TypedefPackage.get_info thy name;
- val rewrite_rule = MetaSimplifier.rewrite_rule [rec_UNIV_I, rec_True_simp];
- in map rewrite_rule [abs_inject, abs_inverse, abs_induct] end;
- val tname = Binding.name (Long_Name.base_name name);
- in
- thy
- |> TypecopyPackage.typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
- |-> (fn (name, _) => `(fn thy => get_thms thy name))
- end;
-
-fun mixit convs refls =
- let fun f ((res,lhs,rhs),refl) = ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs);
- in #1 (Library.foldl f (([],[],convs),refls)) end;
-
-
-fun extension_definition full name fields names alphas zeta moreT more vars thy =
- let
- val base = Long_Name.base_name;
- val fieldTs = (map snd fields);
- val alphas_zeta = alphas@[zeta];
- val alphas_zetaTs = map (fn n => TFree (n, HOLogic.typeS)) alphas_zeta;
- val vT = TFree (Name.variant alphas_zeta "'v", HOLogic.typeS);
- val extT_name = suffix ext_typeN name
- val extT = Type (extT_name, alphas_zetaTs);
- val repT = foldr1 HOLogic.mk_prodT (fieldTs@[moreT]);
- val fields_more = fields@[(full moreN,moreT)];
- val fields_moreTs = fieldTs@[moreT];
- val bfields_more = map (apfst base) fields_more;
- val r = Free (rN,extT)
- val len = length fields;
- val idxms = 0 upto len;
-
- (* prepare declarations and definitions *)
-
- (*fields constructor*)
- val ext_decl = (mk_extC (name,extT) fields_moreTs);
- (*
- val ext_spec = Const ext_decl :==
- (foldr (uncurry lambda)
- (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more]))) (vars@[more]))
- *)
- val ext_spec = list_comb (Const ext_decl,vars@[more]) :==
- (mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more])));
-
- fun mk_ext args = list_comb (Const ext_decl, args);
-
- (*destructors*)
- val _ = timing_msg "record extension preparing definitions";
- val dest_decls = map (mk_selC extT o (apfst (suffix ext_dest))) bfields_more;
-
- fun mk_dest_spec (i, (c,T)) =
- let val snds = (funpow i HOLogic.mk_snd (mk_Rep name repT extT $ r))
- in Const (mk_selC extT (suffix ext_dest c,T))
- :== (lambda r (if i=len then snds else HOLogic.mk_fst snds))
- end;
- val dest_specs =
- ListPair.map mk_dest_spec (idxms, fields_more);
-
- (*updates*)
- val upd_decls = map (mk_updC updN extT) bfields_more;
- fun mk_upd_spec (c,T) =
- let
- val args = map (fn (n,nT) => if n=c then Free (base c,T --> T)$
- (mk_sel r (suffix ext_dest n,nT))
- else (mk_sel r (suffix ext_dest n,nT)))
- fields_more;
- in Const (mk_updC updN extT (c,T))$(Free (base c,T --> T))$r
- :== mk_ext args
- end;
- val upd_specs = map mk_upd_spec fields_more;
-
- (* 1st stage: defs_thy *)
- fun mk_defs () =
- thy
- |> extension_typedef name repT (alphas @ [zeta])
- ||> Sign.add_consts_i
- (map (Syntax.no_syn o apfst Binding.name) (apfst base ext_decl :: dest_decls @ upd_decls))
- ||>> PureThy.add_defs false
- (map (Thm.no_attributes o apfst Binding.name) (ext_spec :: dest_specs))
- ||>> PureThy.add_defs false
- (map (Thm.no_attributes o apfst Binding.name) upd_specs)
- |-> (fn args as ((_, dest_defs), upd_defs) =>
- fold Code.add_default_eqn dest_defs
- #> fold Code.add_default_eqn upd_defs
- #> pair args);
- val ((([abs_inject, abs_inverse, abs_induct], ext_def :: dest_defs), upd_defs), defs_thy) =
- timeit_msg "record extension type/selector/update defs:" mk_defs;
-
- (* prepare propositions *)
- val _ = timing_msg "record extension preparing propositions";
- val vars_more = vars@[more];
- val named_vars_more = (names@[full moreN])~~vars_more;
- val variants = map (fn (Free (x,_))=>x) vars_more;
- val ext = mk_ext vars_more;
- val s = Free (rN, extT);
- val w = Free (wN, extT);
- val P = Free (Name.variant variants "P", extT-->HOLogic.boolT);
- val C = Free (Name.variant variants "C", HOLogic.boolT);
-
- val inject_prop =
- let val vars_more' = map (fn (Free (x,T)) => Free (x ^ "'",T)) vars_more;
- in All (map dest_Free (vars_more@vars_more'))
- ((HOLogic.eq_const extT $
- mk_ext vars_more$mk_ext vars_more')
- ===
- foldr1 HOLogic.mk_conj (map HOLogic.mk_eq (vars_more ~~ vars_more')))
- end;
-
- val induct_prop =
- (All (map dest_Free vars_more) (Trueprop (P $ ext)), Trueprop (P $ s));
-
- val cases_prop =
- (All (map dest_Free vars_more)
- (Trueprop (HOLogic.mk_eq (s,ext)) ==> Trueprop C))
- ==> Trueprop C;
-
- (*destructors*)
- val dest_conv_props =
- map (fn (c, x as Free (_,T)) => mk_sel ext (suffix ext_dest c,T) === x) named_vars_more;
-
- (*updates*)
- fun mk_upd_prop (i,(c,T)) =
- let val x' = Free (Name.variant variants (base c ^ "'"),T --> T)
- val args' = nth_map i (K (x'$nth vars_more i)) vars_more
- in mk_upd updN c x' ext === mk_ext args' end;
- val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
-
- val surjective_prop =
- let val args =
- map (fn (c, Free (_,T)) => mk_sel s (suffix ext_dest c,T)) named_vars_more;
- in s === mk_ext args end;
-
- val split_meta_prop =
- let val P = Free (Name.variant variants "P", extT-->Term.propT) in
- Logic.mk_equals
- (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext))
- end;
-
- fun prove stndrd = quick_and_dirty_prove stndrd defs_thy;
- val prove_standard = quick_and_dirty_prove true defs_thy;
- fun prove_simp stndrd simps =
- let val tac = simp_all_tac HOL_ss simps
- in fn prop => prove stndrd [] prop (K tac) end;
-
- fun inject_prf () = (prove_simp true [ext_def,abs_inject,Pair_eq] inject_prop);
- val inject = timeit_msg "record extension inject proof:" inject_prf;
-
- fun induct_prf () =
- let val (assm, concl) = induct_prop
- in prove_standard [assm] concl (fn {prems, ...} =>
- EVERY [try_param_tac rN abs_induct 1,
- simp_tac (HOL_ss addsimps [split_paired_all]) 1,
- resolve_tac (map (rewrite_rule [ext_def]) prems) 1])
- end;
- val induct = timeit_msg "record extension induct proof:" induct_prf;
-
- fun cases_prf_opt () =
- let
- val (_$(Pvar$_)) = concl_of induct;
- val ind = cterm_instantiate
- [(cterm_of defs_thy Pvar, cterm_of defs_thy
- (lambda w (HOLogic.imp$HOLogic.mk_eq(r,w)$C)))]
- induct;
- in standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
-
- fun cases_prf_noopt () =
- prove_standard [] cases_prop (fn _ =>
- EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1,
- try_param_tac rN induct 1,
- rtac impI 1,
- REPEAT (etac allE 1),
- etac mp 1,
- rtac refl 1])
-
- val cases_prf = quick_and_dirty_prf cases_prf_noopt cases_prf_opt;
- val cases = timeit_msg "record extension cases proof:" cases_prf;
-
- fun dest_convs_prf () = map (prove_simp false
- ([ext_def,abs_inverse]@Pair_sel_convs@dest_defs)) dest_conv_props;
- val dest_convs = timeit_msg "record extension dest_convs proof:" dest_convs_prf;
- fun dest_convs_standard_prf () = map standard dest_convs;
-
- val dest_convs_standard =
- timeit_msg "record extension dest_convs_standard proof:" dest_convs_standard_prf;
-
- fun upd_convs_prf_noopt () = map (prove_simp true (dest_convs_standard@upd_defs))
- upd_conv_props;
- fun upd_convs_prf_opt () =
- let
-
- fun mkrefl (c,T) = Thm.reflexive
- (cterm_of defs_thy (Free (Name.variant variants (base c ^ "'"),T-->T)));
- val refls = map mkrefl fields_more;
- val dest_convs' = map mk_meta_eq dest_convs;
- val map_eqs = map (uncurry Thm.combination) (refls ~~ dest_convs');
-
- val constr_refl = Thm.reflexive (cterm_of defs_thy (head_of ext));
-
- fun mkthm (udef,(fld_refl,thms)) =
- let val bdyeq = Library.foldl (uncurry Thm.combination) (constr_refl,thms);
- (* (|N=N (|N=N,M=M,K=K,more=more|)
- M=M (|N=N,M=M,K=K,more=more|)
- K=K'
- more = more (|N=N,M=M,K=K,more=more|) =
- (|N=N,M=M,K=K',more=more|)
- *)
- val (_$(_$v$r)$_) = prop_of udef;
- val (_$(v'$_)$_) = prop_of fld_refl;
- val udef' = cterm_instantiate
- [(cterm_of defs_thy v,cterm_of defs_thy v'),
- (cterm_of defs_thy r,cterm_of defs_thy ext)] udef;
- in standard (Thm.transitive udef' bdyeq) end;
- in map mkthm (rev upd_defs ~~ (mixit dest_convs' map_eqs)) end;
-
- val upd_convs_prf = quick_and_dirty_prf upd_convs_prf_noopt upd_convs_prf_opt;
-
- val upd_convs =
- timeit_msg "record extension upd_convs proof:" upd_convs_prf;
-
- fun surjective_prf () =
- prove_standard [] surjective_prop (fn _ =>
- (EVERY [try_param_tac rN induct 1,
- simp_tac (HOL_basic_ss addsimps dest_convs_standard) 1]));
- val surjective = timeit_msg "record extension surjective proof:" surjective_prf;
-
- fun split_meta_prf () =
- prove_standard [] split_meta_prop (fn _ =>
- EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
- etac meta_allE 1, atac 1,
- rtac (prop_subst OF [surjective]) 1,
- REPEAT (etac meta_allE 1), atac 1]);
- val split_meta = timeit_msg "record extension split_meta proof:" split_meta_prf;
-
-
- val (([inject',induct',cases',surjective',split_meta'],
- [dest_convs',upd_convs']),
- thm_thy) =
- defs_thy
- |> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
- [("ext_inject", inject),
- ("ext_induct", induct),
- ("ext_cases", cases),
- ("ext_surjective", surjective),
- ("ext_split", split_meta)]
- ||>> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
- [("dest_convs", dest_convs_standard), ("upd_convs", upd_convs)]
-
- in (thm_thy,extT,induct',inject',dest_convs',split_meta',upd_convs')
- end;
-
-fun chunks [] [] = []
- | chunks [] xs = [xs]
- | chunks (l::ls) xs = Library.take (l,xs)::chunks ls (Library.drop (l,xs));
-
-fun chop_last [] = error "last: list should not be empty"
- | chop_last [x] = ([],x)
- | chop_last (x::xs) = let val (tl,l) = chop_last xs in (x::tl,l) end;
-
-fun subst_last s [] = error "subst_last: list should not be empty"
- | subst_last s ([x]) = [s]
- | subst_last s (x::xs) = (x::subst_last s xs);
-
-(* mk_recordT builds up the record type from the current extension tpye extT and a list
- * of parent extensions, starting with the root of the record hierarchy
-*)
-fun mk_recordT extT =
- fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT;
-
-
-
-fun obj_to_meta_all thm =
- let
- fun E thm = case (SOME (spec OF [thm]) handle THM _ => NONE) of
- SOME thm' => E thm'
- | NONE => thm;
- val th1 = E thm;
- val th2 = Drule.forall_intr_vars th1;
- in th2 end;
-
-fun meta_to_obj_all thm =
- let
- val thy = Thm.theory_of_thm thm;
- val prop = Thm.prop_of thm;
- val params = Logic.strip_params prop;
- val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl prop);
- val ct = cterm_of thy
- (HOLogic.mk_Trueprop (HOLogic.list_all (params, concl)));
- val thm' = Seq.hd (REPEAT (rtac allI 1) (Thm.trivial ct));
- in
- Thm.implies_elim thm' thm
- end;
-
-
-
-(* record_definition *)
-
-fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy =
- let
- val external_names = NameSpace.external_names (Sign.naming_of thy);
-
- val alphas = map fst args;
- val name = Sign.full_bname thy bname;
- val full = Sign.full_bname_path thy bname;
- val base = Long_Name.base_name;
-
- val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields);
-
- val parent_fields = List.concat (map #fields parents);
- val parent_chunks = map (length o #fields) parents;
- val parent_names = map fst parent_fields;
- val parent_types = map snd parent_fields;
- val parent_fields_len = length parent_fields;
- val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map base parent_names);
- val parent_vars = ListPair.map Free (parent_variants, parent_types);
- val parent_len = length parents;
- val parents_idx = (map #name parents) ~~ (0 upto (parent_len - 1));
-
- val fields = map (apfst full) bfields;
- val names = map fst fields;
- val extN = full bname;
- val types = map snd fields;
- val alphas_fields = List.foldr OldTerm.add_typ_tfree_names [] types;
- val alphas_ext = alphas inter alphas_fields;
- val len = length fields;
- val variants =
- Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) (map fst bfields);
- val vars = ListPair.map Free (variants, types);
- val named_vars = names ~~ vars;
- val idxs = 0 upto (len - 1);
- val idxms = 0 upto len;
-
- val all_fields = parent_fields @ fields;
- val all_names = parent_names @ names;
- val all_types = parent_types @ types;
- val all_len = parent_fields_len + len;
- val all_variants = parent_variants @ variants;
- val all_vars = parent_vars @ vars;
- val all_named_vars = (parent_names ~~ parent_vars) @ named_vars;
-
-
- val zeta = Name.variant alphas "'z";
- val moreT = TFree (zeta, HOLogic.typeS);
- val more = Free (moreN, moreT);
- val full_moreN = full moreN;
- val bfields_more = bfields @ [(moreN,moreT)];
- val fields_more = fields @ [(full_moreN,moreT)];
- val vars_more = vars @ [more];
- val named_vars_more = named_vars @[(full_moreN,more)];
- val all_vars_more = all_vars @ [more];
- val all_named_vars_more = all_named_vars @ [(full_moreN,more)];
-
- (* 1st stage: extension_thy *)
- val (extension_thy,extT,ext_induct,ext_inject,ext_dest_convs,ext_split,u_convs) =
- thy
- |> Sign.add_path bname
- |> extension_definition full extN fields names alphas_ext zeta moreT more vars;
-
- val _ = timing_msg "record preparing definitions";
- val Type extension_scheme = extT;
- val extension_name = unsuffix ext_typeN (fst extension_scheme);
- val extension = let val (n,Ts) = extension_scheme in (n,subst_last HOLogic.unitT Ts) end;
- val extension_names =
- (map ((unsuffix ext_typeN) o fst o #extension) parents) @ [extN];
- val extension_id = Library.foldl (op ^) ("",extension_names);
-
-
- fun rec_schemeT n = mk_recordT (map #extension (prune n parents)) extT;
- val rec_schemeT0 = rec_schemeT 0;
-
- fun recT n =
- let val (c,Ts) = extension
- in mk_recordT (map #extension (prune n parents)) (Type (c,subst_last HOLogic.unitT Ts))
- end;
- val recT0 = recT 0;
-
- fun mk_rec args n =
- let val (args',more) = chop_last args;
- fun mk_ext' (((name,T),args),more) = mk_ext (name,T) (args@[more]);
- fun build Ts =
- List.foldr mk_ext' more (prune n (extension_names ~~ Ts ~~ (chunks parent_chunks args')))
- in
- if more = HOLogic.unit
- then build (map recT (0 upto parent_len))
- else build (map rec_schemeT (0 upto parent_len))
- end;
-
- val r_rec0 = mk_rec all_vars_more 0;
- val r_rec_unit0 = mk_rec (all_vars@[HOLogic.unit]) 0;
-
- fun r n = Free (rN, rec_schemeT n)
- val r0 = r 0;
- fun r_unit n = Free (rN, recT n)
- val r_unit0 = r_unit 0;
- val w = Free (wN, rec_schemeT 0)
-
- (* prepare print translation functions *)
- val field_tr's =
- print_translation (distinct (op =) (maps external_names (full_moreN :: names)));
-
- val adv_ext_tr's =
- let
- val trnames = external_names extN;
- in map (gen_record_tr') trnames end;
-
- val adv_record_type_abbr_tr's =
- let val trnames = external_names (hd extension_names);
- val lastExt = unsuffix ext_typeN (fst extension);
- in map (gen_record_type_abbr_tr' name alphas zeta lastExt rec_schemeT0) trnames
- end;
-
- val adv_record_type_tr's =
- let val trnames = if parent_len > 0 then external_names extN else [];
- (* avoid conflict with adv_record_type_abbr_tr's *)
- in map (gen_record_type_tr') trnames
- end;
-
-
- (* prepare declarations *)
-
- val sel_decls = map (mk_selC rec_schemeT0) bfields_more;
- val upd_decls = map (mk_updC updateN rec_schemeT0) bfields_more;
- val make_decl = (makeN, all_types ---> recT0);
- val fields_decl = (fields_selN, types ---> Type extension);
- val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0);
- val truncate_decl = (truncateN, rec_schemeT0 --> recT0);
-
- (* prepare definitions *)
-
- fun parent_more s =
- if null parents then s
- else mk_sel s (Long_Name.qualify (#name (List.last parents)) moreN, extT);
-
- fun parent_more_upd v s =
- if null parents then v$s
- else let val mp = Long_Name.qualify (#name (List.last parents)) moreN;
- in mk_upd updateN mp v s end;
-
- (*record (scheme) type abbreviation*)
- val recordT_specs =
- [(Binding.name (suffix schemeN bname), alphas @ [zeta], rec_schemeT0, Syntax.NoSyn),
- (Binding.name bname, alphas, recT0, Syntax.NoSyn)];
-
- (*selectors*)
- fun mk_sel_spec (c,T) =
- Const (mk_selC rec_schemeT0 (c,T))
- :== (lambda r0 (Const (mk_selC extT (suffix ext_dest c,T))$parent_more r0));
- val sel_specs = map mk_sel_spec fields_more;
-
- (*updates*)
-
- fun mk_upd_spec (c,T) =
- let
- val new = mk_upd' updN c (Free (base c,T-->T)) extT(*(parent_more r0)*);
- in Const (mk_updC updateN rec_schemeT0 (c,T))$(Free (base c,T-->T))$r0
- :== (parent_more_upd new r0)
- end;
- val upd_specs = map mk_upd_spec fields_more;
-
- (*derived operations*)
- val make_spec = Const (full makeN, all_types ---> recT0) $$ all_vars :==
- mk_rec (all_vars @ [HOLogic.unit]) 0;
- val fields_spec = Const (full fields_selN, types ---> Type extension) $$ vars :==
- mk_rec (all_vars @ [HOLogic.unit]) parent_len;
- val extend_spec =
- Const (full extendN, recT0-->moreT-->rec_schemeT0) $ r_unit0 $ more :==
- mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0;
- val truncate_spec = Const (full truncateN, rec_schemeT0 --> recT0) $ r0 :==
- mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0;
-
- (* 2st stage: defs_thy *)
-
- fun mk_defs () =
- extension_thy
- |> Sign.add_trfuns
- ([],[],field_tr's, [])
- |> Sign.add_advanced_trfuns
- ([],[],adv_ext_tr's @ adv_record_type_tr's @ adv_record_type_abbr_tr's,[])
- |> Sign.parent_path
- |> Sign.add_tyabbrs_i recordT_specs
- |> Sign.add_path bname
- |> Sign.add_consts_i
- (map2 (fn (x, T) => fn mx => (Binding.name x, T, mx))
- sel_decls (field_syntax @ [Syntax.NoSyn]))
- |> (Sign.add_consts_i o map (fn (x, T) => (Binding.name x, T, Syntax.NoSyn)))
- (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
- |> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) sel_specs)
- ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) upd_specs)
- ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name))
- [make_spec, fields_spec, extend_spec, truncate_spec])
- |-> (fn defs as ((sel_defs, upd_defs), derived_defs) =>
- fold Code.add_default_eqn sel_defs
- #> fold Code.add_default_eqn upd_defs
- #> fold Code.add_default_eqn derived_defs
- #> pair defs)
- val (((sel_defs, upd_defs), derived_defs), defs_thy) =
- timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:"
- mk_defs;
-
-
- (* prepare propositions *)
- val _ = timing_msg "record preparing propositions";
- val P = Free (Name.variant all_variants "P", rec_schemeT0-->HOLogic.boolT);
- val C = Free (Name.variant all_variants "C", HOLogic.boolT);
- val P_unit = Free (Name.variant all_variants "P", recT0-->HOLogic.boolT);
-
- (*selectors*)
- val sel_conv_props =
- map (fn (c, x as Free (_,T)) => mk_sel r_rec0 (c,T) === x) named_vars_more;
-
- (*updates*)
- fun mk_upd_prop (i,(c,T)) =
- let val x' = Free (Name.variant all_variants (base c ^ "'"),T-->T);
- val n = parent_fields_len + i;
- val args' = nth_map n (K (x'$nth all_vars_more n)) all_vars_more
- in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end;
- val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
-
- (*induct*)
- val induct_scheme_prop =
- All (map dest_Free all_vars_more) (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0);
- val induct_prop =
- (All (map dest_Free all_vars) (Trueprop (P_unit $ r_rec_unit0)),
- Trueprop (P_unit $ r_unit0));
-
- (*surjective*)
- val surjective_prop =
- let val args = map (fn (c,Free (_,T)) => mk_sel r0 (c,T)) all_named_vars_more
- in r0 === mk_rec args 0 end;
-
- (*cases*)
- val cases_scheme_prop =
- (All (map dest_Free all_vars_more)
- (Trueprop (HOLogic.mk_eq (r0,r_rec0)) ==> Trueprop C))
- ==> Trueprop C;
-
- val cases_prop =
- (All (map dest_Free all_vars)
- (Trueprop (HOLogic.mk_eq (r_unit0,r_rec_unit0)) ==> Trueprop C))
- ==> Trueprop C;
-
- (*split*)
- val split_meta_prop =
- let val P = Free (Name.variant all_variants "P", rec_schemeT0-->Term.propT) in
- Logic.mk_equals
- (All [dest_Free r0] (P $ r0), All (map dest_Free all_vars_more) (P $ r_rec0))
- end;
-
- val split_object_prop =
- let fun ALL vs t = List.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) t vs
- in (ALL [dest_Free r0] (P $ r0)) === (ALL (map dest_Free all_vars_more) (P $ r_rec0))
- end;
-
-
- val split_ex_prop =
- let fun EX vs t = List.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) t vs
- in (EX [dest_Free r0] (P $ r0)) === (EX (map dest_Free all_vars_more) (P $ r_rec0))
- end;
-
- (*equality*)
- val equality_prop =
- let
- val s' = Free (rN ^ "'", rec_schemeT0)
- fun mk_sel_eq (c,Free (_,T)) = mk_sel r0 (c,T) === mk_sel s' (c,T)
- val seleqs = map mk_sel_eq all_named_vars_more
- in All (map dest_Free [r0,s']) (Logic.list_implies (seleqs,r0 === s')) end;
-
- (* 3rd stage: thms_thy *)
-
- fun prove stndrd = quick_and_dirty_prove stndrd defs_thy;
- val prove_standard = quick_and_dirty_prove true defs_thy;
-
- fun prove_simp stndrd ss simps =
- let val tac = simp_all_tac ss simps
- in fn prop => prove stndrd [] prop (K tac) end;
-
- val ss = get_simpset defs_thy;
-
- fun sel_convs_prf () = map (prove_simp false ss
- (sel_defs@ext_dest_convs)) sel_conv_props;
- val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf;
- fun sel_convs_standard_prf () = map standard sel_convs
- val sel_convs_standard =
- timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf;
-
- fun upd_convs_prf () =
- map (prove_simp true ss (upd_defs@u_convs)) upd_conv_props;
-
- val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf;
-
- val parent_induct = if null parents then [] else [#induct (hd (rev parents))];
-
- fun induct_scheme_prf () = prove_standard [] induct_scheme_prop (fn _ =>
- (EVERY [if null parent_induct
- then all_tac else try_param_tac rN (hd parent_induct) 1,
- try_param_tac rN ext_induct 1,
- asm_simp_tac HOL_basic_ss 1]));
- val induct_scheme = timeit_msg "record induct_scheme proof:" induct_scheme_prf;
-
- fun induct_prf () =
- let val (assm, concl) = induct_prop;
- in
- prove_standard [assm] concl (fn {prems, ...} =>
- try_param_tac rN induct_scheme 1
- THEN try_param_tac "more" @{thm unit.induct} 1
- THEN resolve_tac prems 1)
- end;
- val induct = timeit_msg "record induct proof:" induct_prf;
-
- fun surjective_prf () =
- prove_standard [] surjective_prop (fn prems =>
- (EVERY [try_param_tac rN induct_scheme 1,
- simp_tac (ss addsimps sel_convs_standard) 1]))
- val surjective = timeit_msg "record surjective proof:" surjective_prf;
-
- fun cases_scheme_prf_opt () =
- let
- val (_$(Pvar$_)) = concl_of induct_scheme;
- val ind = cterm_instantiate
- [(cterm_of defs_thy Pvar, cterm_of defs_thy
- (lambda w (HOLogic.imp$HOLogic.mk_eq(r0,w)$C)))]
- induct_scheme;
- in standard (ObjectLogic.rulify (mp OF [ind, refl])) end;
-
- fun cases_scheme_prf_noopt () =
- prove_standard [] cases_scheme_prop (fn _ =>
- EVERY [asm_full_simp_tac (HOL_basic_ss addsimps [atomize_all, atomize_imp]) 1,
- try_param_tac rN induct_scheme 1,
- rtac impI 1,
- REPEAT (etac allE 1),
- etac mp 1,
- rtac refl 1])
- val cases_scheme_prf = quick_and_dirty_prf cases_scheme_prf_noopt cases_scheme_prf_opt;
- val cases_scheme = timeit_msg "record cases_scheme proof:" cases_scheme_prf;
-
- fun cases_prf () =
- prove_standard [] cases_prop (fn _ =>
- try_param_tac rN cases_scheme 1
- THEN simp_all_tac HOL_basic_ss [unit_all_eq1]);
- val cases = timeit_msg "record cases proof:" cases_prf;
-
- fun split_meta_prf () =
- prove false [] split_meta_prop (fn _ =>
- EVERY [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
- etac meta_allE 1, atac 1,
- rtac (prop_subst OF [surjective]) 1,
- REPEAT (etac meta_allE 1), atac 1]);
- val split_meta = timeit_msg "record split_meta proof:" split_meta_prf;
- val split_meta_standard = standard split_meta;
-
- fun split_object_prf_opt () =
- let
- val cPI= cterm_of defs_thy (lambda r0 (Trueprop (P$r0)));
- val (_$Abs(_,_,P$_)) = fst (Logic.dest_equals (concl_of split_meta_standard));
- val cP = cterm_of defs_thy P;
- val split_meta' = cterm_instantiate [(cP,cPI)] split_meta_standard;
- val (l,r) = HOLogic.dest_eq (HOLogic.dest_Trueprop split_object_prop);
- val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l);
- val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r);
- val thl = assume cl (*All r. P r*) (* 1 *)
- |> obj_to_meta_all (*!!r. P r*)
- |> equal_elim split_meta' (*!!n m more. P (ext n m more)*)
- |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*)
- |> implies_intr cl (* 1 ==> 2 *)
- val thr = assume cr (*All n m more. P (ext n m more)*)
- |> obj_to_meta_all (*!!n m more. P (ext n m more)*)
- |> equal_elim (symmetric split_meta') (*!!r. P r*)
- |> meta_to_obj_all (*All r. P r*)
- |> implies_intr cr (* 2 ==> 1 *)
- in standard (thr COMP (thl COMP iffI)) end;
-
- fun split_object_prf_noopt () =
- prove_standard [] split_object_prop (fn _ =>
- EVERY [rtac iffI 1,
- REPEAT (rtac allI 1), etac allE 1, atac 1,
- rtac allI 1, rtac induct_scheme 1,REPEAT (etac allE 1),atac 1]);
-
- val split_object_prf = quick_and_dirty_prf split_object_prf_noopt split_object_prf_opt;
- val split_object = timeit_msg "record split_object proof:" split_object_prf;
-
-
- fun split_ex_prf () =
- prove_standard [] split_ex_prop (fn _ =>
- EVERY [rtac iffI 1,
- etac exE 1,
- simp_tac (HOL_basic_ss addsimps [split_meta_standard]) 1,
- ex_inst_tac 1,
- (*REPEAT (rtac exI 1),*)
- atac 1,
- REPEAT (etac exE 1),
- rtac exI 1,
- atac 1]);
- val split_ex = timeit_msg "record split_ex proof:" split_ex_prf;
-
- fun equality_tac thms =
- let val (s'::s::eqs) = rev thms;
- val ss' = ss addsimps (s'::s::sel_convs_standard);
- val eqs' = map (simplify ss') eqs;
- in simp_tac (HOL_basic_ss addsimps (s'::s::eqs')) 1 end;
-
- fun equality_prf () = prove_standard [] equality_prop (fn {context, ...} =>
- fn st => let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in
- st |> (res_inst_tac context [((rN, 0), s)] cases_scheme 1
- THEN res_inst_tac context [((rN, 0), s')] cases_scheme 1
- THEN (METAHYPS equality_tac 1))
- (* simp_all_tac ss (sel_convs) would also work but is less efficient *)
- end);
- val equality = timeit_msg "record equality proof:" equality_prf;
-
- val ((([sel_convs',upd_convs',sel_defs',upd_defs',[split_meta',split_object',split_ex'],derived_defs'],
- [surjective',equality']),[induct_scheme',induct',cases_scheme',cases']), thms_thy) =
- defs_thy
- |> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
- [("select_convs", sel_convs_standard),
- ("update_convs", upd_convs),
- ("select_defs", sel_defs),
- ("update_defs", upd_defs),
- ("splits", [split_meta_standard,split_object,split_ex]),
- ("defs", derived_defs)]
- ||>> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
- [("surjective", surjective),
- ("equality", equality)]
- ||>> (PureThy.add_thms o (map o apfst o apfst) Binding.name)
- [(("induct_scheme", induct_scheme), induct_type_global (suffix schemeN name)),
- (("induct", induct), induct_type_global name),
- (("cases_scheme", cases_scheme), cases_type_global (suffix schemeN name)),
- (("cases", cases), cases_type_global name)];
-
-
- val sel_upd_simps = sel_convs' @ upd_convs';
- val iffs = [ext_inject]
- val final_thy =
- thms_thy
- |> (snd oo PureThy.add_thmss)
- [((Binding.name "simps", sel_upd_simps),
- [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]),
- ((Binding.name "iffs", iffs), [iff_add])]
- |> put_record name (make_record_info args parent fields extension induct_scheme')
- |> put_sel_upd (names @ [full_moreN]) sel_upd_simps
- |> add_record_equalities extension_id equality'
- |> add_extinjects ext_inject
- |> add_extsplit extension_name ext_split
- |> add_record_splits extension_id (split_meta',split_object',split_ex',induct_scheme')
- |> add_extfields extension_name (fields @ [(full_moreN,moreT)])
- |> add_fieldext (extension_name,snd extension) (names @ [full_moreN])
- |> Sign.parent_path;
-
- in final_thy
- end;
-
-
-(* add_record *)
-
-(*we do all preparations and error checks here, deferring the real
- work to record_definition*)
-fun gen_add_record prep_typ prep_raw_parent quiet_mode (params, bname) raw_parent raw_fields thy =
- let
- val _ = Theory.requires thy "Record" "record definitions";
- val _ = if quiet_mode then () else writeln ("Defining record " ^ quote bname ^ " ...");
-
- val ctxt = ProofContext.init thy;
-
-
- (* parents *)
-
- fun prep_inst T = fst (cert_typ ctxt T []);
-
- val parent = Option.map (apfst (map prep_inst) o prep_raw_parent ctxt) raw_parent
- handle ERROR msg => cat_error msg ("The error(s) above in parent record specification");
- val parents = add_parents thy parent [];
-
- val init_env =
- (case parent of
- NONE => []
- | SOME (types, _) => List.foldr OldTerm.add_typ_tfrees [] types);
-
-
- (* fields *)
-
- fun prep_field (c, raw_T, mx) env =
- let val (T, env') = prep_typ ctxt raw_T env handle ERROR msg =>
- cat_error msg ("The error(s) above occured in record field " ^ quote c)
- in ((c, T, mx), env') end;
-
- val (bfields, envir) = fold_map prep_field raw_fields init_env;
- val envir_names = map fst envir;
-
-
- (* args *)
-
- val defaultS = Sign.defaultS thy;
- val args = map (fn x => (x, AList.lookup (op =) envir x |> the_default defaultS)) params;
-
-
- (* errors *)
-
- val name = Sign.full_bname thy bname;
- val err_dup_record =
- if is_none (get_record thy name) then []
- else ["Duplicate definition of record " ^ quote name];
-
- val err_dup_parms =
- (case duplicates (op =) params of
- [] => []
- | dups => ["Duplicate parameter(s) " ^ commas dups]);
-
- val err_extra_frees =
- (case subtract (op =) params envir_names of
- [] => []
- | extras => ["Extra free type variable(s) " ^ commas extras]);
-
- val err_no_fields = if null bfields then ["No fields present"] else [];
-
- val err_dup_fields =
- (case duplicates (op =) (map #1 bfields) of
- [] => []
- | dups => ["Duplicate field(s) " ^ commas_quote dups]);
-
- val err_bad_fields =
- if forall (not_equal moreN o #1) bfields then []
- else ["Illegal field name " ^ quote moreN];
-
- val err_dup_sorts =
- (case duplicates (op =) envir_names of
- [] => []
- | dups => ["Inconsistent sort constraints for " ^ commas dups]);
-
- val errs =
- err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @
- err_dup_fields @ err_bad_fields @ err_dup_sorts;
- in
- if null errs then () else error (cat_lines errs) ;
- thy |> record_definition (args, bname) parent parents bfields
- end
- handle ERROR msg => cat_error msg ("Failed to define record " ^ quote bname);
-
-val add_record = gen_add_record read_typ read_raw_parent;
-val add_record_i = gen_add_record cert_typ (K I);
-
-(* setup theory *)
-
-val setup =
- Sign.add_trfuns ([], parse_translation, [], []) #>
- Sign.add_advanced_trfuns ([], adv_parse_translation, [], []) #>
- Simplifier.map_simpset (fn ss =>
- ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]);
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val record_decl =
- P.type_args -- P.name --
- (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const);
-
-val _ =
- OuterSyntax.command "record" "define extensible record" K.thy_decl
- (record_decl >> (fn (x, (y, z)) => Toplevel.theory (add_record false x y z)));
-
-end;
-
-end;
-
-
-structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage;
-open BasicRecordPackage;
--- a/src/HOL/Tools/refute.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/refute.ML Tue Jun 23 15:48:55 2009 +0100
@@ -73,7 +73,7 @@
val specialize_type : theory -> (string * Term.typ) -> Term.term -> Term.term
val string_of_typ : Term.typ -> string
val typ_of_dtyp :
- DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> DatatypeAux.dtyp
+ Datatype.descr -> (Datatype.dtyp * Term.typ) list -> Datatype.dtyp
-> Term.typ
end; (* signature REFUTE *)
@@ -411,15 +411,6 @@
(* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *)
(* ------------------------------------------------------------------------- *)
-(* ------------------------------------------------------------------------- *)
-(* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type *)
-(* ('Term.typ'), given type parameters for the data type's type *)
-(* arguments *)
-(* ------------------------------------------------------------------------- *)
-
- (* DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list ->
- DatatypeAux.dtyp -> Term.typ *)
-
fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
(* replace a 'DtTFree' variable by the associated type *)
the (AList.lookup (op =) typ_assoc (DatatypeAux.DtTFree a))
@@ -526,7 +517,7 @@
fun is_IDT_constructor thy (s, T) =
(case body_type T of
Type (s', _) =>
- (case DatatypePackage.get_datatype_constrs thy s' of
+ (case Datatype.get_datatype_constrs thy s' of
SOME constrs =>
List.exists (fn (cname, cty) =>
cname = s andalso Sign.typ_instance thy (T, cty)) constrs
@@ -545,7 +536,7 @@
fun is_IDT_recursor thy (s, T) =
let
val rec_names = Symtab.fold (append o #rec_names o snd)
- (DatatypePackage.get_datatypes thy) []
+ (Datatype.get_datatypes thy) []
in
(* I'm not quite sure if checking the name 's' is sufficient, *)
(* or if we should also check the type 'T'. *)
@@ -834,7 +825,7 @@
(* axiomatic type classes *)
| Type ("itself", [T1]) => collect_type_axioms (axs, T1)
| Type (s, Ts) =>
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info => (* inductive datatype *)
(* only collect relevant type axioms for the argument types *)
Library.foldl collect_type_axioms (axs, Ts)
@@ -969,7 +960,7 @@
Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc)
| Type ("prop", []) => acc
| Type (s, Ts) =>
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info => (* inductive datatype *)
let
val index = #index info
@@ -1181,7 +1172,7 @@
(* TODO: no warning needed for /positive/ occurrences of IDTs *)
val maybe_spurious = Library.exists (fn
Type (s, _) =>
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info => (* inductive datatype *)
let
val index = #index info
@@ -1660,10 +1651,6 @@
(* their arguments) of the size of the argument types *)
(* ------------------------------------------------------------------------- *)
- (* theory -> (Term.typ * int) list -> DatatypeAux.descr ->
- (DatatypeAux.dtyp * Term.typ) list ->
- (string * DatatypeAux.dtyp list) list -> int *)
-
fun size_of_dtyp thy typ_sizes descr typ_assoc constructors =
sum (map (fn (_, dtyps) =>
product (map (size_of_type thy (typ_sizes, []) o
@@ -1986,7 +1973,7 @@
val (typs, terms) = model
(* Term.typ -> (interpretation * model * arguments) option *)
fun interpret_term (Type (s, Ts)) =
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info => (* inductive datatype *)
let
(* int option -- only recursive IDTs have an associated depth *)
@@ -2120,7 +2107,7 @@
HOLogic_empty_set) pairss
end
| Type (s, Ts) =>
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info =>
(case AList.lookup (op =) typs T of
SOME 0 =>
@@ -2144,7 +2131,6 @@
(* decrement depth for the IDT 'T' *)
val typs' = (case AList.lookup (op =) typs T of NONE => typs
| SOME n => AList.update (op =) (T, n-1) typs)
- (* Term.term list -> DatatypeAux.dtyp list -> Term.term list *)
fun constructor_terms terms [] = terms
| constructor_terms terms (d::ds) =
let
@@ -2185,7 +2171,7 @@
Const (s, T) =>
(case body_type T of
Type (s', Ts') =>
- (case DatatypePackage.get_datatype thy s' of
+ (case Datatype.get_datatype thy s' of
SOME info => (* body type is an inductive datatype *)
let
val index = #index info
@@ -2241,7 +2227,6 @@
else ()
(* returns an interpretation where everything is mapped to *)
(* an "undefined" element of the datatype *)
- (* DatatypeAux.dtyp list -> interpretation *)
fun make_undef [] =
Leaf (replicate total False)
| make_undef (d::ds) =
@@ -2253,7 +2238,6 @@
Node (replicate size (make_undef ds))
end
(* returns the interpretation for a constructor *)
- (* int * DatatypeAux.dtyp list -> int * interpretation *)
fun make_constr (offset, []) =
if offset<total then
(offset+1, Leaf ((replicate offset False) @ True ::
@@ -2421,9 +2405,6 @@
(* mutually recursive types must have the same type *)
(* parameters, unless the mutual recursion comes from *)
(* indirect recursion *)
- (* (DatatypeAux.dtyp * Term.typ) list ->
- (DatatypeAux.dtyp * Term.typ) list ->
- (DatatypeAux.dtyp * Term.typ) list *)
fun rec_typ_assoc acc [] =
acc
| rec_typ_assoc acc ((d, T)::xs) =
@@ -2458,7 +2439,6 @@
else
raise REFUTE ("IDT_recursion_interpreter",
"different type associations for the same dtyp"))
- (* (DatatypeAux.dtyp * Term.typ) list *)
val typ_assoc = filter
(fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false)
(rec_typ_assoc []
@@ -2613,7 +2593,6 @@
val rec_dtyps_args = List.filter
(DatatypeAux.is_rec_type o fst) (dtyps ~~ args)
(* map those indices to interpretations *)
- (* (DatatypeAux.dtyp * interpretation) list *)
val rec_dtyps_intrs = map (fn (dtyp, arg) =>
let
val dT = typ_of_dtyp descr typ_assoc dtyp
@@ -2625,8 +2604,6 @@
(* takes the dtyp and interpretation of an element, *)
(* and computes the interpretation for the *)
(* corresponding recursive argument *)
- (* DatatypeAux.dtyp -> interpretation ->
- interpretation *)
fun rec_intr (DatatypeAux.DtRec i) (Leaf xs) =
(* recursive argument is "rec_i params elem" *)
compute_array_entry i (find_index_eq True xs)
@@ -2675,7 +2652,7 @@
end
else
NONE (* not a recursion operator of this datatype *)
- ) (DatatypePackage.get_datatypes thy) NONE
+ ) (Datatype.get_datatypes thy) NONE
| _ => (* head of term is not a constant *)
NONE;
@@ -3224,7 +3201,7 @@
fun IDT_printer thy model T intr assignment =
(case T of
Type (s, Ts) =>
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
SOME info => (* inductive datatype *)
let
val (typs, _) = model
@@ -3252,8 +3229,6 @@
(* constructor generates the datatype's element that is given by *)
(* 'element', returns the constructor (as a term) as well as the *)
(* indices of the arguments *)
- (* string * DatatypeAux.dtyp list ->
- (Term.term * int list) option *)
fun get_constr_args (cname, cargs) =
let
val cTerm = Const (cname,
@@ -3281,7 +3256,6 @@
in
Option.map (fn args => (cTerm, cargs, args)) (get_args iC)
end
- (* Term.term * DatatypeAux.dtyp list * int list *)
val (cTerm, cargs, args) =
(* we could speed things up by computing the correct *)
(* constructor directly (rather than testing all *)
--- a/src/HOL/Tools/res_atp.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/res_atp.ML Tue Jun 23 15:48:55 2009 +0100
@@ -9,6 +9,7 @@
val get_relevant : int -> bool -> Proof.context * (thm list * 'a) -> thm list ->
(thm * (string * int)) list
val prepare_clauses : bool -> thm list -> thm list ->
+ (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list ->
(thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> theory ->
ResHolClause.axiom_name vector * (ResHolClause.clause list * ResHolClause.clause list *
ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list)
@@ -476,7 +477,7 @@
| occ _ = false
in occ end;
-fun is_recordtype T = not (null (RecordPackage.dest_recTs T));
+fun is_recordtype T = not (null (Record.dest_recTs T));
(*Unwanted equalities include
(1) those between a variable that does not properly occur in the second operand,
@@ -526,26 +527,30 @@
relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls)
end;
-(* prepare and count clauses before writing *)
-fun prepare_clauses dfg goal_cls chain_ths axcls thy =
+(* prepare for passing to writer,
+ create additional clauses based on the information from extra_cls *)
+fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy =
let
- val white_thms = filter check_named (map ResAxioms.pairname (if null chain_ths then whitelist else chain_ths))
+ val white_thms =
+ filter check_named (map ResAxioms.pairname (if null chain_ths then whitelist else chain_ths))
val white_cls = ResAxioms.cnf_rules_pairs thy white_thms
- val axcls = white_cls @ axcls
- val ccls = subtract_cls goal_cls axcls
+ val extra_cls = white_cls @ extra_cls
+ val ccls = subtract_cls goal_cls extra_cls
val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls
val ccltms = map prop_of ccls
- and axtms = map (prop_of o #1) axcls
+ and axtms = map (prop_of o #1) extra_cls
val subs = tfree_classes_of_terms ccltms
and supers = tvar_classes_of_terms axtms
and tycons = type_consts_of_terms thy (ccltms@axtms)
(*TFrees in conjecture clauses; TVars in axiom clauses*)
- val (clnames, (conjectures, axiom_clauses, helper_clauses)) =
- ResHolClause.prepare_clauses dfg thy (isFO thy goal_cls) ccls axcls []
+ val conjectures = ResHolClause.make_conjecture_clauses dfg thy ccls
+ val (clnames,axiom_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy axcls)
+ val helper_clauses = ResHolClause.get_helper_clauses dfg thy (isFO thy goal_cls) (conjectures, extra_cls, [])
val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers
val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
in
- (Vector.fromList clnames, (conjectures, axiom_clauses, helper_clauses, classrel_clauses, arity_clauses))
+ (Vector.fromList clnames,
+ (conjectures, axiom_clauses, helper_clauses, classrel_clauses, arity_clauses))
end
end;
--- a/src/HOL/Tools/res_hol_clause.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/res_hol_clause.ML Tue Jun 23 15:48:55 2009 +0100
@@ -28,13 +28,18 @@
val strip_comb: combterm -> combterm * combterm list
val literals_of_term: theory -> term -> literal list * typ list
exception TOO_TRIVIAL
- val count_constants: clause list * clause list * clause list * 'a * 'b ->
- int Symtab.table * bool Symtab.table
- val prepare_clauses: bool -> theory -> bool -> thm list -> (thm * (axiom_name * clause_id)) list ->
- string list -> axiom_name list * (clause list * clause list * clause list)
- val tptp_write_file: string -> int Symtab.table * bool Symtab.table ->
+ val make_conjecture_clauses: bool -> theory -> thm list -> clause list (* dfg thy ccls *)
+ val make_axiom_clauses: bool ->
+ theory ->
+ (thm * (axiom_name * clause_id)) list -> (axiom_name * clause) list (* dfg thy axcls *)
+ val get_helper_clauses: bool ->
+ theory ->
+ bool ->
+ clause list * (thm * (axiom_name * clause_id)) list * string list ->
+ clause list
+ val tptp_write_file: string ->
clause list * clause list * clause list * ResClause.classrelClause list * ResClause.arityClause list -> unit
- val dfg_write_file: string -> int Symtab.table * bool Symtab.table ->
+ val dfg_write_file: string ->
clause list * clause list * clause list * ResClause.classrelClause list * ResClause.arityClause list -> unit
end
@@ -403,10 +408,12 @@
fun cnf_helper_thms thy =
ResAxioms.cnf_rules_pairs thy o map ResAxioms.pairname
-fun get_helper_clauses dfg thy isFO (conjectures, axclauses, user_lemmas) =
+fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) =
if isFO then [] (*first-order*)
else
- let val ct0 = List.foldl count_clause init_counters conjectures
+ let
+ val axclauses = map #2 (make_axiom_clauses dfg thy axcls)
+ val ct0 = List.foldl count_clause init_counters conjectures
val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses
fun needed c = valOf (Symtab.lookup ct c) > 0
val IK = if needed "c_COMBI" orelse needed "c_COMBK"
@@ -464,18 +471,10 @@
(* tptp format *)
-fun prepare_clauses dfg thy isFO thms ax_tuples user_lemmas =
+fun tptp_write_file filename clauses =
let
- val conjectures = make_conjecture_clauses dfg thy thms
- val (clnames,axclauses) = ListPair.unzip (make_axiom_clauses dfg thy ax_tuples)
- val helper_clauses = get_helper_clauses dfg thy isFO (conjectures, axclauses, user_lemmas)
- in
- (clnames, (conjectures, axclauses, helper_clauses))
- end
-
-(* write TPTP format to a single file *)
-fun tptp_write_file filename const_counts (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) =
- let
+ val (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) = clauses
+ val const_counts = count_constants clauses
val (tptp_clss,tfree_litss) = ListPair.unzip (map (clause2tptp const_counts) conjectures)
val tfree_clss = map RC.tptp_tfree_clause (List.foldl (op union_string) [] tfree_litss)
val out = TextIO.openOut filename
@@ -492,8 +491,10 @@
(* dfg format *)
-fun dfg_write_file filename const_counts (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) =
+fun dfg_write_file filename clauses =
let
+ val (conjectures, axclauses, helper_clauses, classrel_clauses, arity_clauses) = clauses
+ val const_counts = count_constants clauses
val (dfg_clss, tfree_litss) = ListPair.unzip (map (clause2dfg const_counts) conjectures)
and probname = Path.implode (Path.base (Path.explode filename))
val axstrs = map (#1 o (clause2dfg const_counts)) axclauses
--- a/src/HOL/Tools/specification_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,257 +0,0 @@
-(* Title: HOL/Tools/specification_package.ML
- Author: Sebastian Skalberg, TU Muenchen
-
-Package for defining constants by specification.
-*)
-
-signature SPECIFICATION_PACKAGE =
-sig
- val add_specification: string option -> (bstring * xstring * bool) list ->
- theory * thm -> theory * thm
-end
-
-structure SpecificationPackage: SPECIFICATION_PACKAGE =
-struct
-
-(* actual code *)
-
-local
- fun mk_definitional [] arg = arg
- | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
- case HOLogic.dest_Trueprop (concl_of thm) of
- Const("Ex",_) $ P =>
- let
- val ctype = domain_type (type_of P)
- val cname_full = Sign.intern_const thy cname
- val cdefname = if thname = ""
- then Thm.def_name (Long_Name.base_name cname)
- else thname
- val def_eq = Logic.mk_equals (Const(cname_full,ctype),
- HOLogic.choice_const ctype $ P)
- val (thms, thy') = PureThy.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
- val thm' = [thm,hd thms] MRS @{thm exE_some}
- in
- mk_definitional cos (thy',thm')
- end
- | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
-
- fun mk_axiomatic axname cos arg =
- let
- fun process [] (thy,tm) =
- let
- val (thms, thy') = PureThy.add_axioms [((Binding.name axname, HOLogic.mk_Trueprop tm),[])] thy
- in
- (thy',hd thms)
- end
- | process ((thname,cname,covld)::cos) (thy,tm) =
- case tm of
- Const("Ex",_) $ P =>
- let
- val ctype = domain_type (type_of P)
- val cname_full = Sign.intern_const thy cname
- val cdefname = if thname = ""
- then Thm.def_name (Long_Name.base_name cname)
- else thname
- val co = Const(cname_full,ctype)
- val thy' = Theory.add_finals_i covld [co] thy
- val tm' = case P of
- Abs(_, _, bodt) => subst_bound (co, bodt)
- | _ => P $ co
- in
- process cos (thy',tm')
- end
- | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
- in
- process cos arg
- end
-
-in
-fun proc_exprop axiomatic cos arg =
- case axiomatic of
- SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
- | NONE => mk_definitional cos arg
-end
-
-fun add_specification axiomatic cos arg =
- arg |> apsnd Thm.freezeT
- |> proc_exprop axiomatic cos
- |> apsnd standard
-
-
-(* Collect all intances of constants in term *)
-
-fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms))
- | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms)
- | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
- | collect_consts ( _,tms) = tms
-
-(* Complementing Type.varify... *)
-
-fun unvarify t fmap =
- let
- val fmap' = map Library.swap fmap
- fun unthaw (f as (a, S)) =
- (case AList.lookup (op =) fmap' a of
- NONE => TVar f
- | SOME (b, _) => TFree (b, S))
- in
- map_types (map_type_tvar unthaw) t
- end
-
-(* The syntactic meddling needed to setup add_specification for work *)
-
-fun process_spec axiomatic cos alt_props thy =
- let
- fun zip3 [] [] [] = []
- | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
- | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error"
-
- fun myfoldr f [x] = x
- | myfoldr f (x::xs) = f (x,myfoldr f xs)
- | myfoldr f [] = error "SpecificationPackage.process_spec internal error"
-
- val rew_imps = alt_props |>
- map (ObjectLogic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
- val props' = rew_imps |>
- map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
-
- fun proc_single prop =
- let
- val frees = OldTerm.term_frees prop
- val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
- orelse error "Specificaton: Only free variables of sort 'type' allowed"
- val prop_closed = List.foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
- in
- (prop_closed,frees)
- end
-
- val props'' = map proc_single props'
- val frees = map snd props''
- val prop = myfoldr HOLogic.mk_conj (map fst props'')
- val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
-
- val (vmap, prop_thawed) = Type.varify [] prop
- val thawed_prop_consts = collect_consts (prop_thawed,[])
- val (altcos,overloaded) = Library.split_list cos
- val (names,sconsts) = Library.split_list altcos
- val consts = map (Syntax.read_term_global thy) sconsts
- val _ = not (Library.exists (not o Term.is_Const) consts)
- orelse error "Specification: Non-constant found as parameter"
-
- fun proc_const c =
- let
- val (_, c') = Type.varify [] c
- val (cname,ctyp) = dest_Const c'
- in
- case List.filter (fn t => let val (name,typ) = dest_Const t
- in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
- end) thawed_prop_consts of
- [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
- | [cf] => unvarify cf vmap
- | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
- end
- val proc_consts = map proc_const consts
- fun mk_exist (c,prop) =
- let
- val T = type_of c
- val cname = Long_Name.base_name (fst (dest_Const c))
- val vname = if Syntax.is_identifier cname
- then cname
- else "x"
- in
- HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
- end
- val ex_prop = List.foldr mk_exist prop proc_consts
- val cnames = map (fst o dest_Const) proc_consts
- fun post_process (arg as (thy,thm)) =
- let
- fun inst_all thy (thm,v) =
- let
- val cv = cterm_of thy v
- val cT = ctyp_of_term cv
- val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
- in
- thm RS spec'
- end
- fun remove_alls frees thm =
- Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees)
- fun process_single ((name,atts),rew_imp,frees) args =
- let
- fun undo_imps thm =
- equal_elim (symmetric rew_imp) thm
-
- fun add_final (arg as (thy, thm)) =
- if name = ""
- then arg |> Library.swap
- else (writeln (" " ^ name ^ ": " ^ (Display.string_of_thm thm));
- PureThy.store_thm (Binding.name name, thm) thy)
- in
- args |> apsnd (remove_alls frees)
- |> apsnd undo_imps
- |> apsnd standard
- |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
- |> add_final
- |> Library.swap
- end
-
- fun process_all [proc_arg] args =
- process_single proc_arg args
- | process_all (proc_arg::rest) (thy,thm) =
- let
- val single_th = thm RS conjunct1
- val rest_th = thm RS conjunct2
- val (thy',_) = process_single proc_arg (thy,single_th)
- in
- process_all rest (thy',rest_th)
- end
- | process_all [] _ = error "SpecificationPackage.process_spec internal error"
- val alt_names = map fst alt_props
- val _ = if exists (fn(name,_) => not (name = "")) alt_names
- then writeln "specification"
- else ()
- in
- arg |> apsnd Thm.freezeT
- |> process_all (zip3 alt_names rew_imps frees)
- end
-
- fun after_qed [[thm]] = ProofContext.theory (fn thy =>
- #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
- in
- thy
- |> ProofContext.init
- |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
- end;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
-val opt_overloaded = P.opt_keyword "overloaded";
-
-val specification_decl =
- P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop)
-
-val _ =
- OuterSyntax.command "specification" "define constants by specification" K.thy_goal
- (specification_decl >> (fn (cos,alt_props) =>
- Toplevel.print o (Toplevel.theory_to_proof
- (process_spec NONE cos alt_props))))
-
-val ax_specification_decl =
- P.name --
- (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop))
-
-val _ =
- OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
- (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
- Toplevel.print o (Toplevel.theory_to_proof
- (process_spec (SOME axname) cos alt_props))))
-
-end
-
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/transfer_data.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,242 @@
+(* Title: Tools/transfer.ML
+ Author: Amine Chaieb, University of Cambridge, 2009
+ Jeremy Avigad, Carnegie Mellon University
+*)
+
+signature TRANSFER_DATA =
+sig
+ type data
+ type entry
+ val get: Proof.context -> data
+ val del: attribute
+ val add: attribute
+ val setup: theory -> theory
+end;
+
+structure TransferData (* : TRANSFER_DATA*) =
+struct
+type entry = {inj : thm list , emb : thm list , ret : thm list , cong : thm list, guess : bool, hints : string list};
+type data = simpset * (thm * entry) list;
+
+val eq_key = Thm.eq_thm;
+fun eq_data arg = eq_fst eq_key arg;
+
+structure Data = GenericDataFun
+(
+ type T = data;
+ val empty = (HOL_ss, []);
+ val extend = I;
+ fun merge _ ((ss, e), (ss', e')) =
+ (merge_ss (ss, ss'), AList.merge eq_key (K true) (e, e'));
+);
+
+val get = Data.get o Context.Proof;
+
+fun del_data key = apsnd (remove eq_data (key, []));
+
+val del = Thm.declaration_attribute (Data.map o del_data);
+val add_ss = Thm.declaration_attribute
+ (fn th => Data.map (fn (ss,data) => (ss addsimps [th], data)));
+
+val del_ss = Thm.declaration_attribute
+ (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
+
+val transM_pat = (Thm.dest_arg1 o Thm.dest_arg o cprop_of) @{thm TransferMorphism_def};
+
+fun merge_update eq m (k,v) [] = [(k,v)]
+ | merge_update eq m (k,v) ((k',v')::al) =
+ if eq (k,k') then (k',m (v,v')):: al else (k',v') :: merge_update eq m (k,v) al
+
+fun C f x y = f y x
+
+fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} =
+ HOL_ss addsimps inj addsimps (if injonly then [] else emb@ret) addcongs cg;
+
+fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th =
+ let
+ val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import_thms true (map Drule.mk_term [a0, D0]) ctxt0)
+ val (aT,bT) =
+ let val T = typ_of (ctyp_of_term a)
+ in (Term.range_type T, Term.domain_type T)
+ end
+ val ctxt' = (Variable.declare_term (term_of a) o Variable.declare_term (term_of D) o Variable.declare_thm th) ctxt
+ val ns = filter (fn i => Type.could_unify (snd i, aT) andalso not (fst (fst i) mem_string leave)) (Term.add_vars (prop_of th) [])
+ val (ins, ctxt'') = Variable.variant_fixes (map (fst o fst) ns) ctxt'
+ val cns = map ((cterm_of o ProofContext.theory_of) ctxt'' o Var) ns
+ val cfis = map ((cterm_of o ProofContext.theory_of) ctxt'' o (fn n => Free (n, bT))) ins
+ val cis = map (Thm.capply a) cfis
+ val (hs,ctxt''') = Assumption.add_assumes (map (fn ct => Thm.capply @{cterm "Trueprop"} (Thm.capply D ct)) cfis) ctxt''
+ val th1 = Drule.cterm_instantiate (cns~~ cis) th
+ val th2 = fold (C implies_elim) hs (fold_rev implies_intr (map cprop_of hs) th1)
+ val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e))
+ (fold_rev implies_intr (map cprop_of hs) th2)
+in hd (Variable.export ctxt''' ctxt0 [th3]) end;
+
+local
+fun transfer_ruleh a D leave ctxt th =
+ let val (ss,al) = get ctxt
+ val a0 = cterm_of (ProofContext.theory_of ctxt) a
+ val D0 = cterm_of (ProofContext.theory_of ctxt) D
+ fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th'
+ in if a0 aconvc a' andalso D0 aconvc D' then SOME e else NONE
+ end
+ in case get_first h al of
+ SOME e => basic_transfer_rule false a0 D0 e leave ctxt th
+ | NONE => error "Transfer: corresponding instance not found in context-data"
+ end
+in fun transfer_rule (a,D) leave (gctxt,th) =
+ (gctxt, transfer_ruleh a D leave (Context.proof_of gctxt) th)
+end;
+
+fun splits P [] = []
+ | splits P (xxs as (x::xs)) =
+ let val pss = filter (P x) xxs
+ val qss = filter_out (P x) xxs
+ in if null pss then [qss] else if null qss then [pss] else pss:: splits P qss
+ end
+
+fun all_transfers leave (gctxt,th) =
+ let
+ val ctxt = Context.proof_of gctxt
+ val tys = map snd (Term.add_vars (prop_of th) [])
+ val _ = if null tys then error "transfer: Unable to guess instance" else ()
+ val tyss = splits (curry Type.could_unify) tys
+ val get_ty = typ_of o ctyp_of_term o fst o Thm.dest_binop o Thm.dest_arg o cprop_of
+ val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
+ val insts =
+ map_filter (fn tys =>
+ get_first (fn (k,ss) => if Type.could_unify (hd tys, range_type (get_ty k))
+ then SOME (get_aD k, ss)
+ else NONE) (snd (get ctxt))) tyss
+ val _ = if null insts then error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction" else ()
+ val ths = map (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts
+ val cth = Conjunction.intr_balanced ths
+ in (gctxt, cth)
+ end;
+
+fun transfer_rule_by_hint ls leave (gctxt,th) =
+ let
+ val ctxt = Context.proof_of gctxt
+ val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
+ val insts =
+ map_filter (fn (k,e) => if exists (fn l => l mem_string (#hints e)) ls
+ then SOME (get_aD k, e) else NONE)
+ (snd (get ctxt))
+ val _ = if null insts then error "Transfer: No labels provided are stored in the context" else ()
+ val ths = map (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts
+ val cth = Conjunction.intr_balanced ths
+ in (gctxt, cth)
+ end;
+
+
+fun transferred_attribute ls NONE leave =
+ if null ls then all_transfers leave else transfer_rule_by_hint ls leave
+ | transferred_attribute _ (SOME (a,D)) leave = transfer_rule (a,D) leave
+
+ (* Add data to the context *)
+fun gen_merge_entries {inj = inj0, emb = emb0, ret = ret0, cong = cg0, guess = g0, hints = hints0}
+ ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
+ {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2})
+ =
+ let fun h xs0 xs ys = subtract Thm.eq_thm xs0 (merge Thm.eq_thm (xs,ys)) in
+ {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
+ ret = h ret0 ret1 ret2, cong = h cg0 cg1 cg2, guess = g1 andalso g2,
+ hints = subtract (op = : string*string -> bool) hints0
+ (hints1 union_string hints2)}
+ end;
+
+local
+ val h = curry (merge Thm.eq_thm)
+in
+fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
+ {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) =
+ {inj = h inj1 inj2, emb = h emb1 emb2, ret = h ret1 ret2, cong = h cg1 cg2, guess = g1 andalso g2, hints = hints1 union_string hints2}
+end;
+
+fun add ((inja,injd), (emba,embd), (reta,retd), (cga,cgd), g, (hintsa, hintsd)) =
+ Thm.declaration_attribute (fn key => fn context => context |> Data.map
+ (fn (ss, al) =>
+ let
+ val _ = ((let val _ = Thm.match (transM_pat, (Thm.dest_arg o cprop_of) key)
+ in 0 end)
+ handle MATCH => error "Attribute expected Theorem of the form : TransferMorphism A a B b")
+ val e0 = {inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa}
+ val ed = {inj = injd, emb = embd, ret = retd, cong = cgd, guess = g, hints = hintsd}
+ val entry =
+ if g then
+ let val (a0,D0) = (Thm.dest_binop o Thm.dest_arg o cprop_of) key
+ val ctxt0 = ProofContext.init (Thm.theory_of_thm key)
+ val inj' = if null inja then #inj (case AList.lookup eq_key al key of SOME e => e | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual")
+ else inja
+ val ret' = merge Thm.eq_thm (reta, map (fn th => basic_transfer_rule true a0 D0 {inj = inj', emb = [], ret = [], cong = cga, guess = g, hints = hintsa} [] ctxt0 th RS sym) emba)
+ in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
+ else e0
+ in (ss, merge_update eq_key (gen_merge_entries ed) (key, entry) al)
+ end));
+
+
+
+(* concrete syntax *)
+
+local
+
+fun keyword k = Scan.lift (Args.$$$ k) >> K ()
+fun keywordC k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+
+val congN = "cong"
+val injN = "inj"
+val embedN = "embed"
+val returnN = "return"
+val addN = "add"
+val delN = "del"
+val modeN = "mode"
+val automaticN = "automatic"
+val manualN = "manual"
+val directionN = "direction"
+val labelsN = "labels"
+val leavingN = "leaving"
+
+val any_keyword = keywordC congN || keywordC injN || keywordC embedN || keywordC returnN || keywordC directionN || keywordC modeN || keywordC delN || keywordC labelsN || keywordC leavingN
+
+val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
+val terms = thms >> map Drule.dest_term
+val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd)
+val name = Scan.lift Args.name
+val names = Scan.repeat (Scan.unless any_keyword name)
+fun optional scan = Scan.optional scan []
+fun optional2 scan = Scan.optional scan ([],[])
+
+val mode = keywordC modeN |-- ((Scan.lift (Args.$$$ manualN) >> K false) || (Scan.lift (Args.$$$ automaticN) >> K true))
+val inj = (keywordC injN |-- thms) -- optional (keywordC delN |-- thms)
+val embed = (keywordC embedN |-- thms) -- optional (keywordC delN |-- thms)
+val return = (keywordC returnN |-- thms) -- optional (keywordC delN |-- thms)
+val cong = (keywordC congN |-- thms) -- optional (keywordC delN |-- thms)
+val addscan = Scan.unless any_keyword (keyword addN)
+val labels = (keywordC labelsN |-- names) -- optional (keywordC delN |-- names)
+val entry = Scan.optional mode true -- optional2 inj -- optional2 embed -- optional2 return -- optional2 cong -- optional2 labels
+
+val transf_add = addscan |-- entry
+in
+
+val install_att_syntax =
+ (Scan.lift (Args.$$$ delN >> K del) ||
+ transf_add
+ >> (fn (((((g, inj), embed), ret), cg), hints) => add (inj, embed, ret, cg, g, hints)))
+
+val transferred_att_syntax = (optional names -- Scan.option (keywordC directionN |-- (Args.term -- Args.term)) -- optional (keywordC leavingN |-- names) >> (fn ((hints, aD),leave) => transferred_attribute hints aD leave));
+
+end;
+
+
+(* theory setup *)
+
+
+val setup =
+ Attrib.setup @{binding transfer} install_att_syntax
+ "Installs transfer data" #>
+ Attrib.setup @{binding transfer_simps} (Attrib.add_del add_ss del_ss)
+ "simp rules for transfer" #>
+ Attrib.setup @{binding transferred} transferred_att_syntax
+ "Transfers theorems";
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/typecopy.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,133 @@
+(* Title: HOL/Tools/typecopy.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Introducing copies of types using trivial typedefs; datatype-like abstraction.
+*)
+
+signature TYPECOPY =
+sig
+ type info = { vs: (string * sort) list, constr: string, typ: typ,
+ inject: thm, proj: string * typ, proj_def: thm }
+ val typecopy: binding * string list -> typ -> (binding * binding) option
+ -> theory -> (string * info) * theory
+ val get_info: theory -> string -> info option
+ val interpretation: (string -> theory -> theory) -> theory -> theory
+ val add_default_code: string -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure Typecopy: TYPECOPY =
+struct
+
+(* theory data *)
+
+type info = {
+ vs: (string * sort) list,
+ constr: string,
+ typ: typ,
+ inject: thm,
+ proj: string * typ,
+ proj_def: thm
+};
+
+structure TypecopyData = TheoryDataFun
+(
+ type T = info Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ = Symtab.merge (K true);
+);
+
+val get_info = Symtab.lookup o TypecopyData.get;
+
+
+(* interpretation of type copies *)
+
+structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
+val interpretation = TypecopyInterpretation.interpretation;
+
+
+(* introducing typecopies *)
+
+fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
+ let
+ val ty = Sign.certify_typ thy raw_ty;
+ val vs =
+ AList.make (the_default HOLogic.typeS o AList.lookup (op =) (Term.add_tfreesT ty [])) raw_vs;
+ val tac = Tactic.rtac UNIV_witness 1;
+ fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
+ Rep_name = c_rep, Abs_inject = inject,
+ Abs_inverse = inverse, ... } : Typedef.info ) thy =
+ let
+ val exists_thm =
+ UNIV_I
+ |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT ty_rep))] [];
+ val inject' = inject OF [exists_thm, exists_thm];
+ val proj_def = inverse OF [exists_thm];
+ val info = {
+ vs = vs,
+ constr = c_abs,
+ typ = ty_rep,
+ inject = inject',
+ proj = (c_rep, ty_abs --> ty_rep),
+ proj_def = proj_def
+ };
+ in
+ thy
+ |> (TypecopyData.map o Symtab.update_new) (tyco, info)
+ |> TypecopyInterpretation.data tyco
+ |> pair (tyco, info)
+ end
+ in
+ thy
+ |> Typedef.add_typedef false (SOME raw_tyco) (raw_tyco, map fst vs, NoSyn)
+ (HOLogic.mk_UNIV ty) (Option.map swap constr_proj) tac
+ |-> (fn (tyco, info) => add_info tyco info)
+ end;
+
+
+(* default code setup *)
+
+fun add_default_code tyco thy =
+ let
+ val SOME { constr = constr_name, proj = (proj, _), proj_def = proj_eq, vs = raw_vs,
+ typ = ty_rep, ... } = get_info thy tyco;
+ val SOME { Rep_inject = proj_inject, ... } = Typedef.get_info thy tyco;
+ val constr = (constr_name, Logic.unvarifyT (Sign.the_const_type thy constr_name));
+ val vs = (map dest_TFree o snd o dest_Type) (Type (tyco, map TFree raw_vs));
+ val ty = Type (tyco, map TFree vs);
+ val proj = Const (proj, ty --> ty_rep);
+ val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
+ val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ $ t_x $ t_y;
+ val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y);
+ val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs);
+ fun tac eq_thm = Class.intro_classes_tac []
+ THEN (Simplifier.rewrite_goals_tac
+ (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
+ THEN ALLGOALS (rtac @{thm refl});
+ fun mk_eq_refl thy = @{thm HOL.eq_refl}
+ |> Thm.instantiate
+ ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
+ |> AxClass.unoverload thy;
+ in
+ thy
+ |> Code.add_datatype [constr]
+ |> Code.add_eqn proj_eq
+ |> TheoryTarget.instantiation ([tyco], vs, [HOLogic.class_eq])
+ |> `(fn lthy => Syntax.check_term lthy eq)
+ |-> (fn eq => Specification.definition
+ (NONE, (Attrib.empty_binding, eq)))
+ |-> (fn (_, (_, eq_thm)) =>
+ Class.prove_instantiation_exit_result Morphism.thm
+ (fn _ => fn eq_thm => tac eq_thm) eq_thm)
+ |-> (fn eq_thm => Code.add_eqn eq_thm)
+ |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy)
+ end;
+
+val setup =
+ TypecopyInterpretation.init
+ #> interpretation add_default_code
+
+end;
--- a/src/HOL/Tools/typecopy_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,157 +0,0 @@
-(* Title: HOL/Tools/typecopy_package.ML
- Author: Florian Haftmann, TU Muenchen
-
-Introducing copies of types using trivial typedefs; datatype-like abstraction.
-*)
-
-signature TYPECOPY_PACKAGE =
-sig
- type info = {
- vs: (string * sort) list,
- constr: string,
- typ: typ,
- inject: thm,
- proj: string * typ,
- proj_def: thm
- }
- val typecopy: binding * string list -> typ -> (binding * binding) option
- -> theory -> (string * info) * theory
- val get_typecopies: theory -> string list
- val get_info: theory -> string -> info option
- val interpretation: (string -> theory -> theory) -> theory -> theory
- val add_default_code: string -> theory -> theory
- val print_typecopies: theory -> unit
- val setup: theory -> theory
-end;
-
-structure TypecopyPackage: TYPECOPY_PACKAGE =
-struct
-
-(* theory data *)
-
-type info = {
- vs: (string * sort) list,
- constr: string,
- typ: typ,
- inject: thm,
- proj: string * typ,
- proj_def: thm
-};
-
-structure TypecopyData = TheoryDataFun
-(
- type T = info Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ = Symtab.merge (K true);
-);
-
-fun print_typecopies thy =
- let
- val tab = TypecopyData.get thy;
- fun mk (tyco, { vs, constr, typ, proj = (proj, _), ... } : info) =
- (Pretty.block o Pretty.breaks) [
- Syntax.pretty_typ_global thy (Type (tyco, map TFree vs)),
- Pretty.str "=",
- (Pretty.str o Sign.extern_const thy) constr,
- Syntax.pretty_typ_global thy typ,
- Pretty.block [Pretty.str "(", (Pretty.str o Sign.extern_const thy) proj, Pretty.str ")"]];
- in
- (Pretty.writeln o Pretty.block o Pretty.fbreaks)
- (Pretty.str "type copies:" :: map mk (Symtab.dest tab))
- end;
-
-val get_typecopies = Symtab.keys o TypecopyData.get;
-val get_info = Symtab.lookup o TypecopyData.get;
-
-
-(* interpretation of type copies *)
-
-structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
-val interpretation = TypecopyInterpretation.interpretation;
-
-
-(* declaring typecopies *)
-
-fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
- let
- val ty = Sign.certify_typ thy raw_ty;
- val vs =
- AList.make (the_default HOLogic.typeS o AList.lookup (op =) (Term.add_tfreesT ty [])) raw_vs;
- val tac = Tactic.rtac UNIV_witness 1;
- fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
- Rep_name = c_rep, Abs_inject = inject,
- Abs_inverse = inverse, ... } : TypedefPackage.info ) thy =
- let
- val exists_thm =
- UNIV_I
- |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT ty_rep))] [];
- val inject' = inject OF [exists_thm, exists_thm];
- val proj_def = inverse OF [exists_thm];
- val info = {
- vs = vs,
- constr = c_abs,
- typ = ty_rep,
- inject = inject',
- proj = (c_rep, ty_abs --> ty_rep),
- proj_def = proj_def
- };
- in
- thy
- |> (TypecopyData.map o Symtab.update_new) (tyco, info)
- |> TypecopyInterpretation.data tyco
- |> pair (tyco, info)
- end
- in
- thy
- |> TypedefPackage.add_typedef false (SOME raw_tyco) (raw_tyco, map fst vs, NoSyn)
- (HOLogic.mk_UNIV ty) (Option.map swap constr_proj) tac
- |-> (fn (tyco, info) => add_info tyco info)
- end;
-
-
-(* default code setup *)
-
-fun add_default_code tyco thy =
- let
- val SOME { constr = constr_name, proj = (proj, _), proj_def = proj_eq, vs = raw_vs,
- typ = ty_rep, ... } = get_info thy tyco;
- val SOME { Rep_inject = proj_inject, ... } = TypedefPackage.get_info thy tyco;
- val constr = (constr_name, Logic.unvarifyT (Sign.the_const_type thy constr_name));
- val vs = (map dest_TFree o snd o dest_Type) (Type (tyco, map TFree raw_vs));
- val ty = Type (tyco, map TFree vs);
- val proj = Const (proj, ty --> ty_rep);
- val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
- val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
- $ t_x $ t_y;
- val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y);
- val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs);
- fun tac eq_thm = Class.intro_classes_tac []
- THEN (Simplifier.rewrite_goals_tac
- (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
- THEN ALLGOALS (rtac @{thm refl});
- fun mk_eq_refl thy = @{thm HOL.eq_refl}
- |> Thm.instantiate
- ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
- |> AxClass.unoverload thy;
- in
- thy
- |> Code.add_datatype [constr]
- |> Code.add_eqn proj_eq
- |> TheoryTarget.instantiation ([tyco], vs, [HOLogic.class_eq])
- |> `(fn lthy => Syntax.check_term lthy eq)
- |-> (fn eq => Specification.definition
- (NONE, (Attrib.empty_binding, eq)))
- |-> (fn (_, (_, eq_thm)) =>
- Class.prove_instantiation_exit_result Morphism.thm
- (fn _ => fn eq_thm => tac eq_thm) eq_thm)
- |-> (fn eq_thm => Code.add_eqn eq_thm)
- |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy)
- end;
-
-val setup =
- TypecopyInterpretation.init
- #> interpretation add_default_code
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/typedef.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,270 @@
+(* Title: HOL/Tools/typedef.ML
+ Author: Markus Wenzel and Stefan Berghofer, TU Muenchen
+
+Gordon/HOL-style type definitions: create a new syntactic type
+represented by a non-empty subset.
+*)
+
+signature TYPEDEF =
+sig
+ type info =
+ {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
+ type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
+ Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
+ Rep_induct: thm, Abs_induct: thm}
+ val add_typedef: bool -> binding option -> binding * string list * mixfix ->
+ term -> (binding * binding) option -> tactic -> theory -> (string * info) * theory
+ val typedef: (bool * binding) * (binding * string list * mixfix) * term
+ * (binding * binding) option -> theory -> Proof.state
+ val typedef_cmd: (bool * binding) * (binding * string list * mixfix) * string
+ * (binding * binding) option -> theory -> Proof.state
+ val get_info: theory -> string -> info option
+ val interpretation: (string -> theory -> theory) -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure Typedef: TYPEDEF =
+struct
+
+(** type definitions **)
+
+(* theory data *)
+
+type info =
+ {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
+ type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
+ Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
+ Rep_induct: thm, Abs_induct: thm};
+
+structure TypedefData = TheoryDataFun
+(
+ type T = info Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ tabs : T = Symtab.merge (K true) tabs;
+);
+
+val get_info = Symtab.lookup o TypedefData.get;
+fun put_info name info = TypedefData.map (Symtab.update (name, info));
+
+
+(* prepare_typedef *)
+
+fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
+
+structure TypedefInterpretation = InterpretationFun(type T = string val eq = op =);
+val interpretation = TypedefInterpretation.interpretation;
+
+fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy =
+ let
+ val _ = Theory.requires thy "Typedef" "typedefs";
+ val ctxt = ProofContext.init thy;
+
+ val full = Sign.full_name thy;
+ val full_name = full name;
+ val bname = Binding.name_of name;
+
+ (*rhs*)
+ val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
+ val setT = Term.fastype_of set;
+ val rhs_tfrees = Term.add_tfrees set [];
+ val rhs_tfreesT = Term.add_tfreesT setT [];
+ val oldT = HOLogic.dest_setT setT handle TYPE _ =>
+ error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
+
+ (*lhs*)
+ val defS = Sign.defaultS thy;
+ val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
+ val args_setT = lhs_tfrees
+ |> filter (member (op =) rhs_tfrees andf (not o member (op =) rhs_tfreesT))
+ |> map TFree;
+
+ val tname = Binding.map_name (Syntax.type_name mx) t;
+ val full_tname = full tname;
+ val newT = Type (full_tname, map TFree lhs_tfrees);
+
+ val (Rep_name, Abs_name) =
+ (case opt_morphs of
+ NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
+ | SOME morphs => morphs);
+ val setT' = map Term.itselfT args_setT ---> setT;
+ val setC = Term.list_comb (Const (full_name, setT'), map Logic.mk_type args_setT);
+ val RepC = Const (full Rep_name, newT --> oldT);
+ val AbsC = Const (full Abs_name, oldT --> newT);
+
+ (*inhabitance*)
+ fun mk_inhabited A =
+ HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)));
+ val set' = if def then setC else set;
+ val goal' = mk_inhabited set';
+ val goal = mk_inhabited set;
+ val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Syntax.read_variable bname), setT));
+
+ (*axiomatization*)
+ val typedef_name = Binding.prefix_name "type_definition_" name;
+ val typedefC =
+ Const (@{const_name type_definition},
+ (newT --> oldT) --> (oldT --> newT) --> setT --> HOLogic.boolT);
+ val typedef_prop = Logic.mk_implies (goal', HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set'));
+ val typedef_deps = Term.add_consts set' [];
+
+ (*set definition*)
+ fun add_def theory =
+ if def then
+ theory
+ |> Sign.add_consts_i [(name, setT', NoSyn)]
+ |> PureThy.add_defs false [Thm.no_attributes (apfst (Binding.name)
+ (PrimitiveDefs.mk_defpair (setC, set)))]
+ |-> (fn [th] => pair (SOME th))
+ else (NONE, theory);
+ fun contract_def NONE th = th
+ | contract_def (SOME def_eq) th =
+ let
+ val cert = Thm.cterm_of (Thm.theory_of_thm def_eq);
+ val goal_eq = MetaSimplifier.rewrite true [def_eq] (cert goal');
+ in Drule.standard (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
+
+ fun typedef_result inhabited =
+ ObjectLogic.typedecl (t, vs, mx)
+ #> snd
+ #> Sign.add_consts_i
+ [(Rep_name, newT --> oldT, NoSyn),
+ (Abs_name, oldT --> newT, NoSyn)]
+ #> add_def
+ #-> (fn set_def =>
+ PureThy.add_axioms [((typedef_name, typedef_prop),
+ [Thm.rule_attribute (K (fn cond_axm => contract_def set_def inhabited RS cond_axm))])]
+ ##>> pair set_def)
+ ##> Theory.add_deps "" (dest_Const RepC) typedef_deps
+ ##> Theory.add_deps "" (dest_Const AbsC) typedef_deps
+ #-> (fn ([type_definition], set_def) => fn thy1 =>
+ let
+ fun make th = Drule.standard (th OF [type_definition]);
+ val ([Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
+ Rep_cases, Abs_cases, Rep_induct, Abs_induct], thy2) =
+ thy1
+ |> Sign.add_path (Binding.name_of name)
+ |> PureThy.add_thms
+ [((Rep_name, make @{thm type_definition.Rep}), []),
+ ((Binding.suffix_name "_inverse" Rep_name, make @{thm type_definition.Rep_inverse}), []),
+ ((Binding.suffix_name "_inverse" Abs_name, make @{thm type_definition.Abs_inverse}), []),
+ ((Binding.suffix_name "_inject" Rep_name, make @{thm type_definition.Rep_inject}), []),
+ ((Binding.suffix_name "_inject" Abs_name, make @{thm type_definition.Abs_inject}), []),
+ ((Binding.suffix_name "_cases" Rep_name, make @{thm type_definition.Rep_cases}),
+ [RuleCases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]),
+ ((Binding.suffix_name "_cases" Abs_name, make @{thm type_definition.Abs_cases}),
+ [RuleCases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]),
+ ((Binding.suffix_name "_induct" Rep_name, make @{thm type_definition.Rep_induct}),
+ [RuleCases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]),
+ ((Binding.suffix_name "_induct" Abs_name, make @{thm type_definition.Abs_induct}),
+ [RuleCases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname])]
+ ||> Sign.parent_path;
+ val info = {rep_type = oldT, abs_type = newT,
+ Rep_name = full Rep_name, Abs_name = full Abs_name,
+ inhabited = inhabited, type_definition = type_definition, set_def = set_def,
+ Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse,
+ Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases,
+ Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct};
+ in
+ thy2
+ |> put_info full_tname info
+ |> TypedefInterpretation.data full_tname
+ |> pair (full_tname, info)
+ end);
+
+
+ (* errors *)
+
+ fun show_names pairs = commas_quote (map fst pairs);
+
+ val illegal_vars =
+ if null (Term.add_vars set []) andalso null (Term.add_tvars set []) then []
+ else ["Illegal schematic variable(s) on rhs"];
+
+ val dup_lhs_tfrees =
+ (case duplicates (op =) lhs_tfrees of [] => []
+ | dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
+
+ val extra_rhs_tfrees =
+ (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => []
+ | extras => ["Extra type variables on rhs: " ^ show_names extras]);
+
+ val illegal_frees =
+ (case Term.add_frees set [] of [] => []
+ | xs => ["Illegal variables on rhs: " ^ show_names xs]);
+
+ val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
+ val _ = if null errs then () else error (cat_lines errs);
+
+ (*test theory errors now!*)
+ val test_thy = Theory.copy thy;
+ val _ = test_thy
+ |> typedef_result (setmp quick_and_dirty true (SkipProof.make_thm test_thy) goal);
+
+ in (set, goal, goal_pat, typedef_result) end
+ handle ERROR msg =>
+ cat_error msg ("The error(s) above occurred in typedef " ^ quote (Binding.str_of name));
+
+
+(* add_typedef: tactic interface *)
+
+fun add_typedef def opt_name typ set opt_morphs tac thy =
+ let
+ val name = the_default (#1 typ) opt_name;
+ val (set, goal, _, typedef_result) =
+ prepare_typedef Syntax.check_term def name typ set opt_morphs thy;
+ val inhabited = Goal.prove_global thy [] [] goal (K tac)
+ handle ERROR msg => cat_error msg
+ ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set));
+ in typedef_result inhabited thy end;
+
+
+(* typedef: proof interface *)
+
+local
+
+fun gen_typedef prep_term ((def, name), typ, set, opt_morphs) thy =
+ let
+ val (_, goal, goal_pat, typedef_result) =
+ prepare_typedef prep_term def name typ set opt_morphs thy;
+ fun after_qed [[th]] = ProofContext.theory (snd o typedef_result th);
+ in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] (ProofContext.init thy) end;
+
+in
+
+val typedef = gen_typedef Syntax.check_term;
+val typedef_cmd = gen_typedef Syntax.read_term;
+
+end;
+
+
+
+(** outer syntax **)
+
+local structure P = OuterParse in
+
+val _ = OuterKeyword.keyword "morphisms";
+
+val typedef_decl =
+ Scan.optional (P.$$$ "(" |--
+ ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
+ --| P.$$$ ")") (true, NONE) --
+ (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
+ Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
+
+fun mk_typedef ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
+ typedef_cmd ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name),
+ (t, vs, mx), A, morphs);
+
+val _ =
+ OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)"
+ OuterKeyword.thy_goal
+ (typedef_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_typedef)));
+
+end;
+
+
+val setup = TypedefInterpretation.init;
+
+end;
--- a/src/HOL/Tools/typedef_codegen.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Tools/typedef_codegen.ML Tue Jun 23 15:48:55 2009 +0100
@@ -24,7 +24,7 @@
val id = Codegen.mk_qual_id module (Codegen.get_const_id gr'' s)
in SOME (Codegen.mk_app brack (Codegen.str id) ps, gr'') end;
fun lookup f T =
- (case TypedefPackage.get_info thy (get_name T) of
+ (case Typedef.get_info thy (get_name T) of
NONE => ""
| SOME info => f info);
in
@@ -45,7 +45,7 @@
| mk_tyexpr ps s = Pretty.list "(" (") " ^ s) ps;
fun typedef_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
- (case TypedefPackage.get_info thy s of
+ (case Typedef.get_info thy s of
NONE => NONE
| SOME {abs_type as newT as Type (tname, Us), rep_type = oldT, Abs_name, Rep_name, ...} =>
if is_some (Codegen.get_assoc_type thy tname) then NONE else
--- a/src/HOL/Tools/typedef_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,270 +0,0 @@
-(* Title: HOL/Tools/typedef_package.ML
- Author: Markus Wenzel and Stefan Berghofer, TU Muenchen
-
-Gordon/HOL-style type definitions: create a new syntactic type
-represented by a non-empty subset.
-*)
-
-signature TYPEDEF_PACKAGE =
-sig
- type info =
- {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
- type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
- Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
- Rep_induct: thm, Abs_induct: thm}
- val get_info: theory -> string -> info option
- val add_typedef: bool -> binding option -> binding * string list * mixfix ->
- term -> (binding * binding) option -> tactic -> theory -> (string * info) * theory
- val typedef: (bool * binding) * (binding * string list * mixfix) * term
- * (binding * binding) option -> theory -> Proof.state
- val typedef_cmd: (bool * binding) * (binding * string list * mixfix) * string
- * (binding * binding) option -> theory -> Proof.state
- val interpretation: (string -> theory -> theory) -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure TypedefPackage: TYPEDEF_PACKAGE =
-struct
-
-(** type definitions **)
-
-(* theory data *)
-
-type info =
- {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
- type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
- Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
- Rep_induct: thm, Abs_induct: thm};
-
-structure TypedefData = TheoryDataFun
-(
- type T = info Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ tabs : T = Symtab.merge (K true) tabs;
-);
-
-val get_info = Symtab.lookup o TypedefData.get;
-fun put_info name info = TypedefData.map (Symtab.update (name, info));
-
-
-(* prepare_typedef *)
-
-fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
-
-structure TypedefInterpretation = InterpretationFun(type T = string val eq = op =);
-val interpretation = TypedefInterpretation.interpretation;
-
-fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy =
- let
- val _ = Theory.requires thy "Typedef" "typedefs";
- val ctxt = ProofContext.init thy;
-
- val full = Sign.full_name thy;
- val full_name = full name;
- val bname = Binding.name_of name;
-
- (*rhs*)
- val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
- val setT = Term.fastype_of set;
- val rhs_tfrees = Term.add_tfrees set [];
- val rhs_tfreesT = Term.add_tfreesT setT [];
- val oldT = HOLogic.dest_setT setT handle TYPE _ =>
- error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
-
- (*lhs*)
- val defS = Sign.defaultS thy;
- val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
- val args_setT = lhs_tfrees
- |> filter (member (op =) rhs_tfrees andf (not o member (op =) rhs_tfreesT))
- |> map TFree;
-
- val tname = Binding.map_name (Syntax.type_name mx) t;
- val full_tname = full tname;
- val newT = Type (full_tname, map TFree lhs_tfrees);
-
- val (Rep_name, Abs_name) =
- (case opt_morphs of
- NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
- | SOME morphs => morphs);
- val setT' = map Term.itselfT args_setT ---> setT;
- val setC = Term.list_comb (Const (full_name, setT'), map Logic.mk_type args_setT);
- val RepC = Const (full Rep_name, newT --> oldT);
- val AbsC = Const (full Abs_name, oldT --> newT);
-
- (*inhabitance*)
- fun mk_inhabited A =
- HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)));
- val set' = if def then setC else set;
- val goal' = mk_inhabited set';
- val goal = mk_inhabited set;
- val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Syntax.read_variable bname), setT));
-
- (*axiomatization*)
- val typedef_name = Binding.prefix_name "type_definition_" name;
- val typedefC =
- Const (@{const_name type_definition},
- (newT --> oldT) --> (oldT --> newT) --> setT --> HOLogic.boolT);
- val typedef_prop = Logic.mk_implies (goal', HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set'));
- val typedef_deps = Term.add_consts set' [];
-
- (*set definition*)
- fun add_def theory =
- if def then
- theory
- |> Sign.add_consts_i [(name, setT', NoSyn)]
- |> PureThy.add_defs false [Thm.no_attributes (apfst (Binding.name)
- (PrimitiveDefs.mk_defpair (setC, set)))]
- |-> (fn [th] => pair (SOME th))
- else (NONE, theory);
- fun contract_def NONE th = th
- | contract_def (SOME def_eq) th =
- let
- val cert = Thm.cterm_of (Thm.theory_of_thm def_eq);
- val goal_eq = MetaSimplifier.rewrite true [def_eq] (cert goal');
- in Drule.standard (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
-
- fun typedef_result inhabited =
- ObjectLogic.typedecl (t, vs, mx)
- #> snd
- #> Sign.add_consts_i
- [(Rep_name, newT --> oldT, NoSyn),
- (Abs_name, oldT --> newT, NoSyn)]
- #> add_def
- #-> (fn set_def =>
- PureThy.add_axioms [((typedef_name, typedef_prop),
- [Thm.rule_attribute (K (fn cond_axm => contract_def set_def inhabited RS cond_axm))])]
- ##>> pair set_def)
- ##> Theory.add_deps "" (dest_Const RepC) typedef_deps
- ##> Theory.add_deps "" (dest_Const AbsC) typedef_deps
- #-> (fn ([type_definition], set_def) => fn thy1 =>
- let
- fun make th = Drule.standard (th OF [type_definition]);
- val ([Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
- Rep_cases, Abs_cases, Rep_induct, Abs_induct], thy2) =
- thy1
- |> Sign.add_path (Binding.name_of name)
- |> PureThy.add_thms
- [((Rep_name, make @{thm type_definition.Rep}), []),
- ((Binding.suffix_name "_inverse" Rep_name, make @{thm type_definition.Rep_inverse}), []),
- ((Binding.suffix_name "_inverse" Abs_name, make @{thm type_definition.Abs_inverse}), []),
- ((Binding.suffix_name "_inject" Rep_name, make @{thm type_definition.Rep_inject}), []),
- ((Binding.suffix_name "_inject" Abs_name, make @{thm type_definition.Abs_inject}), []),
- ((Binding.suffix_name "_cases" Rep_name, make @{thm type_definition.Rep_cases}),
- [RuleCases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]),
- ((Binding.suffix_name "_cases" Abs_name, make @{thm type_definition.Abs_cases}),
- [RuleCases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]),
- ((Binding.suffix_name "_induct" Rep_name, make @{thm type_definition.Rep_induct}),
- [RuleCases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]),
- ((Binding.suffix_name "_induct" Abs_name, make @{thm type_definition.Abs_induct}),
- [RuleCases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname])]
- ||> Sign.parent_path;
- val info = {rep_type = oldT, abs_type = newT,
- Rep_name = full Rep_name, Abs_name = full Abs_name,
- inhabited = inhabited, type_definition = type_definition, set_def = set_def,
- Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse,
- Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases,
- Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct};
- in
- thy2
- |> put_info full_tname info
- |> TypedefInterpretation.data full_tname
- |> pair (full_tname, info)
- end);
-
-
- (* errors *)
-
- fun show_names pairs = commas_quote (map fst pairs);
-
- val illegal_vars =
- if null (Term.add_vars set []) andalso null (Term.add_tvars set []) then []
- else ["Illegal schematic variable(s) on rhs"];
-
- val dup_lhs_tfrees =
- (case duplicates (op =) lhs_tfrees of [] => []
- | dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
-
- val extra_rhs_tfrees =
- (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => []
- | extras => ["Extra type variables on rhs: " ^ show_names extras]);
-
- val illegal_frees =
- (case Term.add_frees set [] of [] => []
- | xs => ["Illegal variables on rhs: " ^ show_names xs]);
-
- val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
- val _ = if null errs then () else error (cat_lines errs);
-
- (*test theory errors now!*)
- val test_thy = Theory.copy thy;
- val _ = test_thy
- |> typedef_result (setmp quick_and_dirty true (SkipProof.make_thm test_thy) goal);
-
- in (set, goal, goal_pat, typedef_result) end
- handle ERROR msg =>
- cat_error msg ("The error(s) above occurred in typedef " ^ quote (Binding.str_of name));
-
-
-(* add_typedef: tactic interface *)
-
-fun add_typedef def opt_name typ set opt_morphs tac thy =
- let
- val name = the_default (#1 typ) opt_name;
- val (set, goal, _, typedef_result) =
- prepare_typedef Syntax.check_term def name typ set opt_morphs thy;
- val inhabited = Goal.prove_global thy [] [] goal (K tac)
- handle ERROR msg => cat_error msg
- ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set));
- in typedef_result inhabited thy end;
-
-
-(* typedef: proof interface *)
-
-local
-
-fun gen_typedef prep_term ((def, name), typ, set, opt_morphs) thy =
- let
- val (_, goal, goal_pat, typedef_result) =
- prepare_typedef prep_term def name typ set opt_morphs thy;
- fun after_qed [[th]] = ProofContext.theory (snd o typedef_result th);
- in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] (ProofContext.init thy) end;
-
-in
-
-val typedef = gen_typedef Syntax.check_term;
-val typedef_cmd = gen_typedef Syntax.read_term;
-
-end;
-
-
-
-(** outer syntax **)
-
-local structure P = OuterParse in
-
-val _ = OuterKeyword.keyword "morphisms";
-
-val typedef_decl =
- Scan.optional (P.$$$ "(" |--
- ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
- --| P.$$$ ")") (true, NONE) --
- (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
-
-fun mk_typedef ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
- typedef_cmd ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name),
- (t, vs, mx), A, morphs);
-
-val _ =
- OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)"
- OuterKeyword.thy_goal
- (typedef_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_typedef)));
-
-end;
-
-
-val setup = TypedefInterpretation.init;
-
-end;
--- a/src/HOL/Typedef.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Typedef.thy Tue Jun 23 15:48:55 2009 +0100
@@ -7,8 +7,8 @@
theory Typedef
imports Set
uses
- ("Tools/typedef_package.ML")
- ("Tools/typecopy_package.ML")
+ ("Tools/typedef.ML")
+ ("Tools/typecopy.ML")
("Tools/typedef_codegen.ML")
begin
@@ -115,8 +115,8 @@
end
-use "Tools/typedef_package.ML" setup TypedefPackage.setup
-use "Tools/typecopy_package.ML" setup TypecopyPackage.setup
+use "Tools/typedef.ML" setup Typedef.setup
+use "Tools/typecopy.ML" setup Typecopy.setup
use "Tools/typedef_codegen.ML" setup TypedefCodegen.setup
end
--- a/src/HOL/Typerep.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Typerep.thy Tue Jun 23 15:48:55 2009 +0100
@@ -64,7 +64,7 @@
in
add_typerep @{type_name fun}
-#> TypedefPackage.interpretation ensure_typerep
+#> Typedef.interpretation ensure_typerep
#> Code.type_interpretation (ensure_typerep o fst)
end
--- a/src/HOL/Wellfounded.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/Wellfounded.thy Tue Jun 23 15:48:55 2009 +0100
@@ -8,7 +8,7 @@
theory Wellfounded
imports Finite_Set Transitive_Closure
-uses ("Tools/function_package/size.ML")
+uses ("Tools/Function/size.ML")
begin
subsection {* Basic Definitions *}
@@ -693,7 +693,7 @@
lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
by (auto simp:inv_image_def)
-text {* Measure functions into @{typ nat} *}
+text {* Measure Datatypes into @{typ nat} *}
definition measure :: "('a => nat) => ('a * 'a)set"
where "measure == inv_image less_than"
@@ -733,7 +733,7 @@
"[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)"
by (unfold trans_def lex_prod_def, blast)
-text {* lexicographic combinations with measure functions *}
+text {* lexicographic combinations with measure Datatypes *}
definition
mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
@@ -948,7 +948,7 @@
subsection {* size of a datatype value *}
-use "Tools/function_package/size.ML"
+use "Tools/Function/size.ML"
setup Size.setup
--- a/src/HOL/ex/Codegenerator_Candidates.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/Codegenerator_Candidates.thy Tue Jun 23 15:48:55 2009 +0100
@@ -25,20 +25,6 @@
"~~/src/HOL/ex/Records"
begin
-(*temporary Haskell workaround*)
-declare typerep_fun_def [code inline]
-declare typerep_prod_def [code inline]
-declare typerep_sum_def [code inline]
-declare typerep_cpoint_ext_type_def [code inline]
-declare typerep_itself_def [code inline]
-declare typerep_list_def [code inline]
-declare typerep_option_def [code inline]
-declare typerep_point_ext_type_def [code inline]
-declare typerep_point'_ext_type_def [code inline]
-declare typerep_point''_ext_type_def [code inline]
-declare typerep_pol_def [code inline]
-declare typerep_polex_def [code inline]
-
(*avoid popular infixes*)
code_reserved SML union inter upto
--- a/src/HOL/ex/LocaleTest2.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/LocaleTest2.thy Tue Jun 23 15:48:55 2009 +0100
@@ -599,7 +599,7 @@
apply (rule_tac x = "gcd x y" in exI)
apply auto [1]
apply (rule_tac x = "lcm x y" in exI)
- apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
+ apply (auto intro: nat_lcm_dvd1 nat_lcm_dvd2 nat_lcm_least)
done
then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" .
txt {* Interpretation to ease use of definitions, which are
@@ -613,7 +613,7 @@
apply (unfold nat_dvd.join_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_sup_def)
- by (auto intro: lcm_dvd1 lcm_dvd2 lcm_least)
+ by (auto intro: nat_lcm_dvd1 nat_lcm_dvd2 nat_lcm_least)
qed
text {* Interpreted theorems from the locales, involving defined terms. *}
@@ -622,7 +622,7 @@
lemma "((x::nat) dvd y & x ~= y) = (x dvd y & x ~= y)"
apply (rule nat_dvd.less_def) done
thm nat_dvd.meet_left text {* from dlat *}
-lemma "gcd x y dvd x"
+lemma "gcd x y dvd (x::nat)"
apply (rule nat_dvd.meet_left) done
--- a/src/HOL/ex/Records.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/Records.thy Tue Jun 23 15:48:55 2009 +0100
@@ -247,8 +247,8 @@
text {* In some cases its convenient to automatically split
(quantified) records. For this purpose there is the simproc @{ML [source]
-"RecordPackage.record_split_simproc"} and the tactic @{ML [source]
-"RecordPackage.record_split_simp_tac"}. The simplification procedure
+"Record.record_split_simproc"} and the tactic @{ML [source]
+"Record.record_split_simp_tac"}. The simplification procedure
only splits the records, whereas the tactic also simplifies the
resulting goal with the standard record simplification rules. A
(generalized) predicate on the record is passed as parameter that
@@ -257,51 +257,51 @@
the quantifier). The value @{ML "0"} indicates no split, a value
greater @{ML "0"} splits up to the given bound of record extension and
finally the value @{ML "~1"} completely splits the record.
-@{ML [source] "RecordPackage.record_split_simp_tac"} additionally takes a list of
+@{ML [source] "Record.record_split_simp_tac"} additionally takes a list of
equations for simplification and can also split fixed record variables.
*}
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
apply (tactic {* simp_tac
- (HOL_basic_ss addsimprocs [RecordPackage.record_split_simproc (K ~1)]) 1*})
+ (HOL_basic_ss addsimprocs [Record.record_split_simproc (K ~1)]) 1*})
apply simp
done
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply simp
done
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
apply (tactic {* simp_tac
- (HOL_basic_ss addsimprocs [RecordPackage.record_split_simproc (K ~1)]) 1*})
+ (HOL_basic_ss addsimprocs [Record.record_split_simproc (K ~1)]) 1*})
apply simp
done
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply simp
done
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
apply (tactic {* simp_tac
- (HOL_basic_ss addsimprocs [RecordPackage.record_split_simproc (K ~1)]) 1*})
+ (HOL_basic_ss addsimprocs [Record.record_split_simproc (K ~1)]) 1*})
apply auto
done
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply auto
done
lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply auto
done
lemma fixes r shows "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply auto
done
@@ -314,7 +314,7 @@
have "\<exists>x. P x"
using pre
apply -
- apply (tactic {* RecordPackage.record_split_simp_tac [] (K ~1) 1*})
+ apply (tactic {* Record.record_split_simp_tac [] (K ~1) 1*})
apply auto
done
}
@@ -322,13 +322,13 @@
qed
text {* The effect of simproc @{ML [source]
-"RecordPackage.record_ex_sel_eq_simproc"} is illustrated by the
+"Record.record_ex_sel_eq_simproc"} is illustrated by the
following lemma.
*}
lemma "\<exists>r. xpos r = x"
apply (tactic {*simp_tac
- (HOL_basic_ss addsimprocs [RecordPackage.record_ex_sel_eq_simproc]) 1*})
+ (HOL_basic_ss addsimprocs [Record.record_ex_sel_eq_simproc]) 1*})
done
--- a/src/HOL/ex/Sqrt.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/Sqrt.thy Tue Jun 23 15:48:55 2009 +0100
@@ -5,7 +5,7 @@
header {* Square roots of primes are irrational *}
theory Sqrt
-imports Complex_Main Primes
+imports Complex_Main
begin
text {*
@@ -14,12 +14,12 @@
*}
theorem sqrt_prime_irrational:
- assumes "prime p"
+ assumes "prime (p::nat)"
shows "sqrt (real p) \<notin> \<rat>"
proof
- from `prime p` have p: "1 < p" by (simp add: prime_def)
+ from `prime p` have p: "1 < p" by (simp add: prime_nat_def)
assume "sqrt (real p) \<in> \<rat>"
- then obtain m n where
+ then obtain m n :: nat where
n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
and gcd: "gcd m n = 1" by (rule Rats_abs_nat_div_natE)
have eq: "m\<twosuperior> = p * n\<twosuperior>"
@@ -34,12 +34,12 @@
have "p dvd m \<and> p dvd n"
proof
from eq have "p dvd m\<twosuperior>" ..
- with `prime p` show "p dvd m" by (rule prime_dvd_power_two)
+ with `prime p` pos2 show "p dvd m" by (rule nat_prime_dvd_power)
then obtain k where "m = p * k" ..
with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
then have "p dvd n\<twosuperior>" ..
- with `prime p` show "p dvd n" by (rule prime_dvd_power_two)
+ with `prime p` pos2 show "p dvd n" by (rule nat_prime_dvd_power)
qed
then have "p dvd gcd m n" ..
with gcd have "p dvd 1" by simp
@@ -48,7 +48,7 @@
qed
corollary "sqrt (real (2::nat)) \<notin> \<rat>"
- by (rule sqrt_prime_irrational) (rule two_is_prime)
+ by (rule sqrt_prime_irrational) (rule nat_two_is_prime)
subsection {* Variations *}
@@ -60,12 +60,12 @@
*}
theorem
- assumes "prime p"
+ assumes "prime (p::nat)"
shows "sqrt (real p) \<notin> \<rat>"
proof
- from `prime p` have p: "1 < p" by (simp add: prime_def)
+ from `prime p` have p: "1 < p" by (simp add: prime_nat_def)
assume "sqrt (real p) \<in> \<rat>"
- then obtain m n where
+ then obtain m n :: nat where
n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
and gcd: "gcd m n = 1" by (rule Rats_abs_nat_div_natE)
from n and sqrt_rat have "real m = \<bar>sqrt (real p)\<bar> * real n" by simp
@@ -75,13 +75,13 @@
also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
then have "p dvd m\<twosuperior>" ..
- with `prime p` have dvd_m: "p dvd m" by (rule prime_dvd_power_two)
+ with `prime p` pos2 have dvd_m: "p dvd m" by (rule nat_prime_dvd_power)
then obtain k where "m = p * k" ..
with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
then have "p dvd n\<twosuperior>" ..
- with `prime p` have "p dvd n" by (rule prime_dvd_power_two)
- with dvd_m have "p dvd gcd m n" by (rule gcd_greatest)
+ with `prime p` pos2 have "p dvd n" by (rule nat_prime_dvd_power)
+ with dvd_m have "p dvd gcd m n" by (rule nat_gcd_greatest)
with gcd have "p dvd 1" by simp
then have "p \<le> 1" by (simp add: dvd_imp_le)
with p show False by simp
--- a/src/HOL/ex/Tarski.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/Tarski.thy Tue Jun 23 15:48:55 2009 +0100
@@ -824,11 +824,6 @@
apply (simp add: intY1_def interval_def intY1_elem)
done
-lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
-apply (rule restrictI)
-apply (erule intY1_f_closed)
-done
-
lemma (in Tarski) intY1_mono:
"monotone (%x: intY1. f x) intY1 (induced intY1 r)"
apply (auto simp add: monotone_def induced_def intY1_f_closed)
@@ -853,7 +848,7 @@
apply (rule CLF.glbH_is_fixp [OF CLF.intro, unfolded CLF_set_def, of "\<lparr>pset = intY1, order = induced intY1 r\<rparr>", simplified])
apply auto
apply (rule intY1_is_cl)
-apply (rule intY1_func)
+apply (erule intY1_f_closed)
apply (rule intY1_mono)
done
--- a/src/HOL/ex/predicate_compile.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOL/ex/predicate_compile.ML Tue Jun 23 15:48:55 2009 +0100
@@ -50,7 +50,7 @@
(* reference to preprocessing of InductiveSet package *)
-val ind_set_codegen_preproc = InductiveSetPackage.codegen_preproc;
+val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
(** fundamentals **)
@@ -279,7 +279,7 @@
end;
fun fetch_pred_data thy name =
- case try (InductivePackage.the_inductive (ProofContext.init thy)) name of
+ case try (Inductive.the_inductive (ProofContext.init thy)) name of
SOME (info as (_, result)) =>
let
fun is_intro_of intro =
@@ -288,7 +288,7 @@
in (fst (dest_Const const) = name) end;
val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result))
val elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
- val nparams = length (InductivePackage.params_of (#raw_induct result))
+ val nparams = length (Inductive.params_of (#raw_induct result))
in (intros, elim, nparams) end
| NONE => error ("No such predicate: " ^ quote name)
@@ -333,7 +333,7 @@
let
val cnstrs = flat (maps
(map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
- (Symtab.dest (DatatypePackage.get_datatypes thy)));
+ (Symtab.dest (Datatype.get_datatypes thy)));
fun check t = (case strip_comb t of
(Free _, []) => true
| (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
@@ -568,7 +568,7 @@
val v = Free (name, T);
val v' = Free (name', T);
in
- lambda v (fst (DatatypePackage.make_case
+ lambda v (fst (Datatype.make_case
(ProofContext.init thy) false [] v
[(mk_tuple out_ts,
if null eqs'' then success_t
@@ -875,7 +875,7 @@
(* else false *)
fun is_constructor thy t =
if (is_Type (fastype_of t)) then
- (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of
+ (case Datatype.get_datatype thy ((fst o dest_Type o fastype_of) t) of
NONE => false
| SOME info => (let
val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
@@ -954,7 +954,7 @@
fun prove_match thy (out_ts : term list) = let
fun get_case_rewrite t =
if (is_constructor thy t) then let
- val case_rewrites = (#case_rewrites (DatatypePackage.the_datatype thy
+ val case_rewrites = (#case_rewrites (Datatype.the_datatype thy
((fst o dest_Type o fastype_of) t)))
in case_rewrites @ (flat (map get_case_rewrite (snd (strip_comb t)))) end
else []
@@ -1067,7 +1067,7 @@
| select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
fun prove_one_direction thy all_vs param_vs modes clauses ((pred, T), mode) = let
-(* val ind_result = InductivePackage.the_inductive (ProofContext.init thy) pred
+(* val ind_result = Inductive.the_inductive (ProofContext.init thy) pred
val index = find_index (fn s => s = pred) (#names (fst ind_result))
val (_, T) = dest_Const (nth (#preds (snd ind_result)) index) *)
val nargs = length (binder_types T) - nparams_of thy pred
@@ -1093,7 +1093,7 @@
fun split_term_tac (Free _) = all_tac
| split_term_tac t =
if (is_constructor thy t) then let
- val info = DatatypePackage.the_datatype thy ((fst o dest_Type o fastype_of) t)
+ val info = Datatype.the_datatype thy ((fst o dest_Type o fastype_of) t)
val num_of_constrs = length (#case_rewrites info)
(* special treatment of pairs -- because of fishing *)
val split_rules = case (fst o dest_Type o fastype_of) t of
@@ -1414,7 +1414,7 @@
fun dependencies_of thy name =
let
fun is_inductive_predicate thy name =
- is_some (try (InductivePackage.the_inductive (ProofContext.init thy)) name)
+ is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
val (intro, elim, nparams) = fetch_pred_data thy name
val data = mk_pred_data ((intro, SOME elim, nparams), [])
val intros = map Thm.prop_of (#intros (rep_pred_data data))
--- a/src/HOLCF/Fixrec.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/Fixrec.thy Tue Jun 23 15:48:55 2009 +0100
@@ -6,7 +6,7 @@
theory Fixrec
imports Sprod Ssum Up One Tr Fix
-uses ("Tools/fixrec_package.ML")
+uses ("Tools/fixrec.ML")
begin
subsection {* Maybe monad type *}
@@ -599,12 +599,12 @@
subsection {* Initializing the fixrec package *}
-use "Tools/fixrec_package.ML"
+use "Tools/fixrec.ML"
-setup {* FixrecPackage.setup *}
+setup {* Fixrec.setup *}
setup {*
- FixrecPackage.add_matchers
+ Fixrec.add_matchers
[ (@{const_name up}, @{const_name match_up}),
(@{const_name sinl}, @{const_name match_sinl}),
(@{const_name sinr}, @{const_name match_sinr}),
--- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy Tue Jun 23 15:48:55 2009 +0100
@@ -184,7 +184,7 @@
val subgoal = Thm.term_of csubgoal;
in
(let
- val weak_case_congs = (map (#weak_case_cong o snd) o Symtab.dest o DatatypePackage.get_datatypes) sign;
+ val weak_case_congs = (map (#weak_case_cong o snd) o Symtab.dest o Datatype.get_datatypes) sign;
val concl = Logic.strip_imp_concl subgoal;
val ic_str = delete_ul_string(Syntax.string_of_term_global sign (IntC sign concl));
val ia_str = delete_ul_string(Syntax.string_of_term_global sign (IntA sign concl));
--- a/src/HOLCF/IOA/meta_theory/Abstraction.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/IOA/meta_theory/Abstraction.thy Tue Jun 23 15:48:55 2009 +0100
@@ -6,7 +6,7 @@
theory Abstraction
imports LiveIOA
-uses ("ioa_package.ML")
+uses ("automaton.ML")
begin
defaultsort type
@@ -613,6 +613,6 @@
method_setup abstraction = {* Scan.succeed (SIMPLE_METHOD' o abstraction_tac) *} ""
-use "ioa_package.ML"
+use "automaton.ML"
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/IOA/meta_theory/automaton.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,536 @@
+(* Title: HOLCF/IOA/meta_theory/automaton.ML
+ Author: Tobias Hamberger, TU Muenchen
+*)
+
+signature AUTOMATON =
+sig
+ val add_ioa: string -> string
+ -> (string) list -> (string) list -> (string) list
+ -> (string * string) list -> string
+ -> (string * string * (string * string)list) list
+ -> theory -> theory
+ val add_composition : string -> (string)list -> theory -> theory
+ val add_hiding : string -> string -> (string)list -> theory -> theory
+ val add_restriction : string -> string -> (string)list -> theory -> theory
+ val add_rename : string -> string -> string -> theory -> theory
+end;
+
+structure Automaton: AUTOMATON =
+struct
+
+val string_of_typ = PrintMode.setmp [] o Syntax.string_of_typ_global;
+val string_of_term = PrintMode.setmp [] o Syntax.string_of_term_global;
+
+exception malformed;
+
+(* stripping quotes *)
+fun strip [] = [] |
+strip ("\""::r) = strip r |
+strip (a::r) = a :: (strip r);
+fun strip_quote s = implode(strip(explode(s)));
+
+(* used by *_of_varlist *)
+fun extract_first (a,b) = strip_quote a;
+fun extract_second (a,b) = strip_quote b;
+(* following functions producing sth from a varlist *)
+fun comma_list_of_varlist [] = "" |
+comma_list_of_varlist [a] = extract_first a |
+comma_list_of_varlist (a::r) = (extract_first a) ^ "," ^ (comma_list_of_varlist r);
+fun primed_comma_list_of_varlist [] = "" |
+primed_comma_list_of_varlist [a] = (extract_first a) ^ "'" |
+primed_comma_list_of_varlist (a::r) = (extract_first a) ^ "'," ^
+ (primed_comma_list_of_varlist r);
+fun type_product_of_varlist [] = "" |
+type_product_of_varlist [a] = "(" ^ extract_second(a) ^ ")" |
+type_product_of_varlist(a::r) = "(" ^ extract_second(a) ^ ")*" ^ type_product_of_varlist r;
+
+(* listing a list *)
+fun list_elements_of [] = "" |
+list_elements_of (a::r) = a ^ " " ^ (list_elements_of r);
+
+(* extracting type parameters from a type list *)
+(* fun param_tupel thy [] res = res |
+param_tupel thy ((Type(_,l))::r) res = param_tupel thy (l @ r) res |
+param_tupel thy ((TFree(a,_))::r) res =
+if (a mem res) then (param_tupel thy r res) else (param_tupel thy r (a::res)) |
+param_tupel thy (a::r) res =
+error ("one component of a statetype is a TVar: " ^ (string_of_typ thy a));
+*)
+
+(* used by constr_list *)
+fun extract_constrs thy [] = [] |
+extract_constrs thy (a::r) =
+let
+fun delete_bold [] = []
+| delete_bold (x::xs) = if (is_prefix (op =) ("\^["::"["::"1"::"m"::[]) (x::xs))
+ then (let val (_::_::_::s) = xs in delete_bold s end)
+ else (if (is_prefix (op =) ("\^["::"["::"0"::"m"::[]) (x::xs))
+ then (let val (_::_::_::s) = xs in delete_bold s end)
+ else (x::delete_bold xs));
+fun delete_bold_string s = implode(delete_bold(explode s));
+(* from a constructor term in *.induct (with quantifiers,
+"Trueprop" and ?P stripped away) delivers the head *)
+fun extract_hd (_ $ Abs(_,_,r)) = extract_hd r |
+extract_hd (Const("Trueprop",_) $ r) = extract_hd r |
+extract_hd (Var(_,_) $ r) = extract_hd r |
+extract_hd (a $ b) = extract_hd a |
+extract_hd (Const(s,_)) = s |
+extract_hd _ = raise malformed;
+(* delivers constructor term string from a prem of *.induct *)
+fun extract_constr thy (_ $ Abs(a,T,r)) = extract_constr thy (snd(Syntax.variant_abs(a,T,r)))|
+extract_constr thy (Const("Trueprop",_) $ r) = extract_constr thy r |
+extract_constr thy (Var(_,_) $ r) = delete_bold_string(string_of_term thy r) |
+extract_constr _ _ = raise malformed;
+in
+(extract_hd a,extract_constr thy a) :: (extract_constrs thy r)
+end;
+
+(* delivering list of constructor terms of a datatype *)
+fun constr_list thy atyp =
+let
+fun act_name thy (Type(s,_)) = s |
+act_name _ s =
+error("malformed action type: " ^ (string_of_typ thy s));
+fun afpl ("." :: a) = [] |
+afpl [] = [] |
+afpl (a::r) = a :: (afpl r);
+fun unqualify s = implode(rev(afpl(rev(explode s))));
+val q_atypstr = act_name thy atyp;
+val uq_atypstr = unqualify q_atypstr;
+val prem = prems_of (PureThy.get_thm thy (uq_atypstr ^ ".induct"));
+in
+extract_constrs thy prem
+handle malformed =>
+error("malformed theorem : " ^ uq_atypstr ^ ".induct")
+end;
+
+fun check_for_constr thy atyp (a $ b) =
+let
+fun all_free (Free(_,_)) = true |
+all_free (a $ b) = if (all_free a) then (all_free b) else false |
+all_free _ = false;
+in
+if (all_free b) then (check_for_constr thy atyp a) else false
+end |
+check_for_constr thy atyp (Const(a,_)) =
+let
+val cl = constr_list thy atyp;
+fun fstmem a [] = false |
+fstmem a ((f,s)::r) = if (a=f) then true else (fstmem a r)
+in
+if (fstmem a cl) then true else false
+end |
+check_for_constr _ _ _ = false;
+
+(* delivering the free variables of a constructor term *)
+fun free_vars_of (t1 $ t2) = (free_vars_of t1) @ (free_vars_of t2) |
+free_vars_of (Const(_,_)) = [] |
+free_vars_of (Free(a,_)) = [a] |
+free_vars_of _ = raise malformed;
+
+(* making a constructor set from a constructor term (of signature) *)
+fun constr_set_string thy atyp ctstr =
+let
+val trm = OldGoals.simple_read_term thy atyp ctstr;
+val l = free_vars_of trm
+in
+if (check_for_constr thy atyp trm) then
+(if (l=[]) then ("{" ^ ctstr ^ "}")
+else "(UN " ^ (list_elements_of l) ^ ". {" ^ ctstr ^ "})")
+else (raise malformed)
+handle malformed =>
+error("malformed action term: " ^ (string_of_term thy trm))
+end;
+
+(* extracting constructor heads *)
+fun constructor_head thy atypstr s =
+let
+fun hd_of (Const(a,_)) = a |
+hd_of (t $ _) = hd_of t |
+hd_of _ = raise malformed;
+val trm = OldGoals.simple_read_term thy (Syntax.read_typ_global thy atypstr) s;
+in
+hd_of trm handle malformed =>
+error("malformed constructor of datatype " ^ atypstr ^ ": " ^ s)
+end;
+fun constructor_head_list _ _ [] = [] |
+constructor_head_list thy atypstr (a::r) =
+ (constructor_head thy atypstr a)::(constructor_head_list thy atypstr r);
+
+(* producing an action set *)
+(*FIXME*)
+fun action_set_string thy atyp [] = "Set.empty" |
+action_set_string thy atyp [a] = constr_set_string thy atyp (strip_quote a) |
+action_set_string thy atyp (a::r) = (constr_set_string thy atyp (strip_quote a)) ^
+ " Un " ^ (action_set_string thy atyp r);
+
+(* used by extend *)
+fun pstr s [] = "(" ^ s ^ "' = " ^ s ^ ")" |
+pstr s ((a,b)::r) =
+if (s=a) then ("(" ^ s ^ "' = (" ^ b ^ "))") else (pstr s r);
+fun poststring [] l = "" |
+poststring [(a,b)] l = pstr a l |
+poststring ((a,b)::r) l = (pstr a l) ^ " & " ^ (poststring r l);
+
+(* extends a (action string,condition,assignlist) tupel by a
+(action term,action string,condition,pseudo_condition,bool) tupel, used by extended_list
+(where bool indicates whether there is a precondition *)
+fun extend thy atyp statetupel (actstr,r,[]) =
+let
+val trm = OldGoals.simple_read_term thy atyp actstr;
+val rtrm = OldGoals.simple_read_term thy (Type("bool",[])) r;
+val flag = if (rtrm=Const("True",Type("bool",[]))) then false else true
+in
+if (check_for_constr thy atyp trm)
+then (trm,actstr, "(" ^ r ^ ") & " ^ (poststring statetupel []),r,flag)
+else
+error("transition " ^ actstr ^ " is not a pure constructor term")
+end |
+extend thy atyp statetupel (actstr,r,(a,b)::c) =
+let
+fun pseudo_poststring [] = "" |
+pseudo_poststring ((a,b)::[]) = "(" ^ a ^ " = (" ^ b ^ "))" |
+pseudo_poststring ((a,b)::r) = "(" ^ a ^ " = (" ^ b ^ ")) & " ^ (pseudo_poststring r);
+val trm = OldGoals.simple_read_term thy atyp actstr;
+val rtrm = OldGoals.simple_read_term thy (Type("bool",[])) r;
+val flag = if (rtrm=Const("True",Type("bool",[]))) then false else true
+in
+if (check_for_constr thy atyp trm) then
+(if ((a="") andalso (b="") andalso (c=[])) then (trm,actstr,r,"True",false)
+(* the case with transrel *)
+ else
+ (trm,actstr,"(" ^ r ^ ") & " ^ (poststring statetupel ((a,b)::c)),
+ "(" ^ r ^ ") & " ^ (pseudo_poststring ((a,b)::c)),flag))
+else
+error("transition " ^ actstr ^ " is not a pure constructor term")
+end;
+(* used by make_alt_string *)
+fun extended_list _ _ _ [] = [] |
+extended_list thy atyp statetupel (a::r) =
+ (extend thy atyp statetupel a) :: (extended_list thy atyp statetupel r);
+
+(* used by write_alts *)
+fun write_alt thy (chead,tr) inp out int [] =
+if (chead mem inp) then
+(
+error("Input action " ^ tr ^ " was not specified")
+) else (
+if (chead mem (out@int)) then
+(writeln("Action " ^ tr ^ " was completedly disabled due to missing specification")) else ();
+(tr ^ " => False",tr ^ " => False")) |
+write_alt thy (chead,ctrm) inp out int ((a,b,c,d,e)::r) =
+let
+fun hd_of (Const(a,_)) = a |
+hd_of (t $ _) = hd_of t |
+hd_of _ = raise malformed;
+fun occurs_again c [] = false |
+occurs_again c ((a,_,_,_,_)::r) = if (c=(hd_of a)) then true else (occurs_again c r);
+in
+if (chead=(hd_of a)) then
+(if ((chead mem inp) andalso e) then (
+error("Input action " ^ b ^ " has a precondition")
+) else (if (chead mem (inp@out@int)) then
+ (if (occurs_again chead r) then (
+error("Two specifications for action: " ^ b)
+ ) else (b ^ " => " ^ c,b ^ " => " ^ d))
+ else (
+error("Action " ^ b ^ " is not in automaton signature")
+))) else (write_alt thy (chead,ctrm) inp out int r)
+handle malformed =>
+error ("malformed action term: " ^ (string_of_term thy a))
+end;
+
+(* used by make_alt_string *)
+fun write_alts thy (a,b) inp out int [] ttr = (a,b) |
+write_alts thy (a,b) inp out int [c] ttr =
+let
+val wa = write_alt thy c inp out int ttr
+in
+ (a ^ (fst wa),b ^ (snd wa))
+end |
+write_alts thy (a,b) inp out int (c::r) ttr =
+let
+val wa = write_alt thy c inp out int ttr
+in
+ write_alts thy (a ^ (fst wa) ^ " | ", b ^ (snd wa) ^ " | ") inp out int r ttr
+end;
+
+fun make_alt_string thy inp out int atyp statetupel trans =
+let
+val cl = constr_list thy atyp;
+val ttr = extended_list thy atyp statetupel trans;
+in
+write_alts thy ("","") inp out int cl ttr
+end;
+
+(* used in add_ioa *)
+fun check_free_primed (Free(a,_)) =
+let
+val (f::r) = rev(explode a)
+in
+if (f="'") then [a] else []
+end |
+check_free_primed (a $ b) = ((check_free_primed a) @ (check_free_primed b)) |
+check_free_primed (Abs(_,_,t)) = check_free_primed t |
+check_free_primed _ = [];
+
+fun overlap [] _ = true |
+overlap (a::r) l = if (a mem l) then (
+error("Two occurences of action " ^ a ^ " in automaton signature")
+) else (overlap r l);
+
+(* delivering some types of an automaton *)
+fun aut_type_of thy aut_name =
+let
+fun left_of (( _ $ left) $ _) = left |
+left_of _ = raise malformed;
+val aut_def = concl_of (PureThy.get_thm thy (aut_name ^ "_def"));
+in
+(#T(rep_cterm(cterm_of thy (left_of aut_def))))
+handle malformed => error ("malformed_theorem : " ^ aut_name ^ "_def")
+end;
+
+fun act_type_of thy (Type(pair1,(Type(pair_asig,(Type(actionset,[acttyp])::_))::_))) = acttyp |
+act_type_of thy t = error ("could not extract action type of following automaton type:\n" ^
+(string_of_typ thy t));
+fun st_type_of thy (Type(pair1,_::(Type(pair2,Type(initial_set,[statetyp])::_))::_)) = statetyp |
+st_type_of thy t = error ("could not extract state type of following automaton type:\n" ^
+(string_of_typ thy t));
+
+fun comp_st_type_of thy [a] = st_type_of thy (aut_type_of thy a) |
+comp_st_type_of thy (a::r) = Type("*",[st_type_of thy (aut_type_of thy a), comp_st_type_of thy r]) |
+comp_st_type_of _ _ = error "empty automaton list";
+
+(* checking consistency of action types (for composition) *)
+fun check_ac thy (a::r) =
+let
+fun ch_f_a thy acttyp [] = acttyp |
+ch_f_a thy acttyp (a::r) =
+let
+val auttyp = aut_type_of thy a;
+val ac = (act_type_of thy auttyp);
+in
+if (ac=acttyp) then (ch_f_a thy acttyp r) else (error "A")
+end;
+val auttyp = aut_type_of thy a;
+val acttyp = (act_type_of thy auttyp);
+in
+ch_f_a thy acttyp r
+end |
+check_ac _ [] = error "empty automaton list";
+
+fun clist [] = "" |
+clist [a] = a |
+clist (a::r) = a ^ " || " ^ (clist r);
+
+val add_defs = snd oo (PureThy.add_defs_cmd false o map Thm.no_attributes o map (apfst Binding.name));
+
+
+(* add_ioa *)
+
+fun add_ioa automaton_name action_type inp out int statetupel ini trans thy =
+(writeln("Constructing automaton " ^ automaton_name ^ " ...");
+let
+val state_type_string = type_product_of_varlist(statetupel);
+val styp = Syntax.read_typ_global thy state_type_string;
+val state_vars_tupel = "(" ^ (comma_list_of_varlist statetupel) ^ ")";
+val state_vars_primed = "(" ^ (primed_comma_list_of_varlist statetupel) ^ ")";
+val atyp = Syntax.read_typ_global thy action_type;
+val inp_set_string = action_set_string thy atyp inp;
+val out_set_string = action_set_string thy atyp out;
+val int_set_string = action_set_string thy atyp int;
+val inp_head_list = constructor_head_list thy action_type inp;
+val out_head_list = constructor_head_list thy action_type out;
+val int_head_list = constructor_head_list thy action_type int;
+val overlap_flag = ((overlap inp out) andalso (overlap inp int) andalso (overlap out int));
+val alt_string = make_alt_string thy inp_head_list out_head_list int_head_list
+ atyp statetupel trans;
+val thy2 = (thy
+|> Sign.add_consts
+[(Binding.name (automaton_name ^ "_initial"), "(" ^ state_type_string ^ ")set" ,NoSyn),
+(Binding.name (automaton_name ^ "_asig"), "(" ^ action_type ^ ")signature" ,NoSyn),
+(Binding.name (automaton_name ^ "_trans"),
+ "(" ^ action_type ^ "," ^ state_type_string ^ ")transition set" ,NoSyn),
+(Binding.name automaton_name, "(" ^ action_type ^ "," ^ state_type_string ^ ")ioa" ,NoSyn)]
+|> add_defs
+[(automaton_name ^ "_initial_def",
+automaton_name ^ "_initial == {" ^ state_vars_tupel ^ "." ^ ini ^ "}"),
+(automaton_name ^ "_asig_def",
+automaton_name ^ "_asig == (" ^
+ inp_set_string ^ "," ^ out_set_string ^ "," ^ int_set_string ^ ")"),
+(automaton_name ^ "_trans_def",
+automaton_name ^ "_trans == {(" ^
+ state_vars_tupel ^ ", act_of_" ^ automaton_name ^ ", " ^ state_vars_primed ^
+"). case act_of_" ^ automaton_name ^ " of " ^ fst(alt_string) ^ "}"),
+(automaton_name ^ "_def",
+automaton_name ^ " == (" ^ automaton_name ^ "_asig, " ^ automaton_name ^
+"_initial, " ^ automaton_name ^ "_trans,{},{})")
+])
+val chk_prime_list = (check_free_primed (OldGoals.simple_read_term thy2 (Type("bool",[]))
+( "case act_of_" ^ automaton_name ^ " of " ^ snd(alt_string))))
+in
+(
+if (chk_prime_list = []) then thy2
+else (
+error("Precondition or assignment terms in postconditions contain following primed variables:\n"
+ ^ (list_elements_of chk_prime_list)))
+)
+end)
+
+fun add_composition automaton_name aut_list thy =
+(writeln("Constructing automaton " ^ automaton_name ^ " ...");
+let
+val acttyp = check_ac thy aut_list;
+val st_typ = comp_st_type_of thy aut_list;
+val comp_list = clist aut_list;
+in
+thy
+|> Sign.add_consts_i
+[(Binding.name automaton_name,
+Type("*",
+[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
+ Type("*",[Type("set",[st_typ]),
+ Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
+ Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
+,NoSyn)]
+|> add_defs
+[(automaton_name ^ "_def",
+automaton_name ^ " == " ^ comp_list)]
+end)
+
+fun add_restriction automaton_name aut_source actlist thy =
+(writeln("Constructing automaton " ^ automaton_name ^ " ...");
+let
+val auttyp = aut_type_of thy aut_source;
+val acttyp = act_type_of thy auttyp;
+val rest_set = action_set_string thy acttyp actlist
+in
+thy
+|> Sign.add_consts_i
+[(Binding.name automaton_name, auttyp,NoSyn)]
+|> add_defs
+[(automaton_name ^ "_def",
+automaton_name ^ " == restrict " ^ aut_source ^ " " ^ rest_set)]
+end)
+fun add_hiding automaton_name aut_source actlist thy =
+(writeln("Constructing automaton " ^ automaton_name ^ " ...");
+let
+val auttyp = aut_type_of thy aut_source;
+val acttyp = act_type_of thy auttyp;
+val hid_set = action_set_string thy acttyp actlist
+in
+thy
+|> Sign.add_consts_i
+[(Binding.name automaton_name, auttyp,NoSyn)]
+|> add_defs
+[(automaton_name ^ "_def",
+automaton_name ^ " == hide " ^ aut_source ^ " " ^ hid_set)]
+end)
+
+fun ren_act_type_of thy funct =
+ (case Term.fastype_of (Syntax.read_term_global thy funct) of
+ Type ("fun", [a, b]) => a
+ | _ => error "could not extract argument type of renaming function term");
+
+fun add_rename automaton_name aut_source fun_name thy =
+(writeln("Constructing automaton " ^ automaton_name ^ " ...");
+let
+val auttyp = aut_type_of thy aut_source;
+val st_typ = st_type_of thy auttyp;
+val acttyp = ren_act_type_of thy fun_name
+in
+thy
+|> Sign.add_consts_i
+[(Binding.name automaton_name,
+Type("*",
+[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
+ Type("*",[Type("set",[st_typ]),
+ Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
+ Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
+,NoSyn)]
+|> add_defs
+[(automaton_name ^ "_def",
+automaton_name ^ " == rename " ^ aut_source ^ " (" ^ fun_name ^ ")")]
+end)
+
+
+(** outer syntax **)
+
+(* prepare results *)
+
+(*encoding transition specifications with a element of ParseTrans*)
+datatype ParseTrans = Rel of string | PP of string*(string*string)list;
+fun mk_trans_of_rel s = Rel(s);
+fun mk_trans_of_prepost (s,l) = PP(s,l);
+
+fun trans_of (a, Rel b) = (a, b, [("", "")])
+ | trans_of (a, PP (b, l)) = (a, b, l);
+
+
+fun mk_ioa_decl (aut, ((((((action_type, inp), out), int), states), initial), trans)) =
+ add_ioa aut action_type inp out int states initial (map trans_of trans);
+
+fun mk_composition_decl (aut, autlist) =
+ add_composition aut autlist;
+
+fun mk_hiding_decl (aut, (actlist, source_aut)) =
+ add_hiding aut source_aut actlist;
+
+fun mk_restriction_decl (aut, (source_aut, actlist)) =
+ add_restriction aut source_aut actlist;
+
+fun mk_rename_decl (aut, (source_aut, rename_f)) =
+ add_rename aut source_aut rename_f;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = List.app OuterKeyword.keyword ["signature", "actions", "inputs",
+ "outputs", "internals", "states", "initially", "transitions", "pre",
+ "post", "transrel", ":=", "compose", "hide_action", "in", "restrict", "to",
+ "rename"];
+
+val actionlist = P.list1 P.term;
+val inputslist = P.$$$ "inputs" |-- P.!!! actionlist;
+val outputslist = P.$$$ "outputs" |-- P.!!! actionlist;
+val internalslist = P.$$$ "internals" |-- P.!!! actionlist;
+val stateslist = P.$$$ "states" |-- P.!!! (Scan.repeat1 (P.name --| P.$$$ "::" -- P.typ));
+val initial = P.$$$ "initially" |-- P.!!! P.term;
+val assign_list = P.list1 (P.name --| P.$$$ ":=" -- P.!!! P.term);
+val pre = P.$$$ "pre" |-- P.!!! P.term;
+val post = P.$$$ "post" |-- P.!!! assign_list;
+val pre1 = (pre -- (Scan.optional post [])) >> mk_trans_of_prepost;
+val post1 = ((Scan.optional pre "True") -- post) >> mk_trans_of_prepost;
+val transrel = (P.$$$ "transrel" |-- P.!!! P.term) >> mk_trans_of_rel;
+val transition = P.term -- (transrel || pre1 || post1);
+val translist = P.$$$ "transitions" |-- P.!!! (Scan.repeat1 transition);
+
+val ioa_decl =
+ (P.name -- (P.$$$ "=" |--
+ (P.$$$ "signature" |--
+ P.!!! (P.$$$ "actions" |--
+ (P.typ --
+ (Scan.optional inputslist []) --
+ (Scan.optional outputslist []) --
+ (Scan.optional internalslist []) --
+ stateslist --
+ (Scan.optional initial "True") --
+ translist))))) >> mk_ioa_decl ||
+ (P.name -- (P.$$$ "=" |-- (P.$$$ "compose" |-- P.!!! (P.list1 P.name))))
+ >> mk_composition_decl ||
+ (P.name -- (P.$$$ "=" |-- (P.$$$ "hide_action" |--
+ P.!!! (P.list1 P.term -- (P.$$$ "in" |-- P.name))))) >> mk_hiding_decl ||
+ (P.name -- (P.$$$ "=" |-- (P.$$$ "restrict" |--
+ P.!!! (P.name -- (P.$$$ "to" |-- P.list1 P.term))))) >> mk_restriction_decl ||
+ (P.name -- (P.$$$ "=" |-- (P.$$$ "rename" |-- (P.!!! (P.name -- (P.$$$ "to" |-- P.term))))))
+ >> mk_rename_decl;
+
+val _ =
+ OuterSyntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" K.thy_decl
+ (ioa_decl >> Toplevel.theory);
+
+end;
+
+end;
--- a/src/HOLCF/IOA/meta_theory/ioa_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,537 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/ioa_package.ML
- ID: $Id$
- Author: Tobias Hamberger, TU Muenchen
-*)
-
-signature IOA_PACKAGE =
-sig
- val add_ioa: string -> string
- -> (string) list -> (string) list -> (string) list
- -> (string * string) list -> string
- -> (string * string * (string * string)list) list
- -> theory -> theory
- val add_composition : string -> (string)list -> theory -> theory
- val add_hiding : string -> string -> (string)list -> theory -> theory
- val add_restriction : string -> string -> (string)list -> theory -> theory
- val add_rename : string -> string -> string -> theory -> theory
-end;
-
-structure IoaPackage: IOA_PACKAGE =
-struct
-
-val string_of_typ = PrintMode.setmp [] o Syntax.string_of_typ_global;
-val string_of_term = PrintMode.setmp [] o Syntax.string_of_term_global;
-
-exception malformed;
-
-(* stripping quotes *)
-fun strip [] = [] |
-strip ("\""::r) = strip r |
-strip (a::r) = a :: (strip r);
-fun strip_quote s = implode(strip(explode(s)));
-
-(* used by *_of_varlist *)
-fun extract_first (a,b) = strip_quote a;
-fun extract_second (a,b) = strip_quote b;
-(* following functions producing sth from a varlist *)
-fun comma_list_of_varlist [] = "" |
-comma_list_of_varlist [a] = extract_first a |
-comma_list_of_varlist (a::r) = (extract_first a) ^ "," ^ (comma_list_of_varlist r);
-fun primed_comma_list_of_varlist [] = "" |
-primed_comma_list_of_varlist [a] = (extract_first a) ^ "'" |
-primed_comma_list_of_varlist (a::r) = (extract_first a) ^ "'," ^
- (primed_comma_list_of_varlist r);
-fun type_product_of_varlist [] = "" |
-type_product_of_varlist [a] = "(" ^ extract_second(a) ^ ")" |
-type_product_of_varlist(a::r) = "(" ^ extract_second(a) ^ ")*" ^ type_product_of_varlist r;
-
-(* listing a list *)
-fun list_elements_of [] = "" |
-list_elements_of (a::r) = a ^ " " ^ (list_elements_of r);
-
-(* extracting type parameters from a type list *)
-(* fun param_tupel thy [] res = res |
-param_tupel thy ((Type(_,l))::r) res = param_tupel thy (l @ r) res |
-param_tupel thy ((TFree(a,_))::r) res =
-if (a mem res) then (param_tupel thy r res) else (param_tupel thy r (a::res)) |
-param_tupel thy (a::r) res =
-error ("one component of a statetype is a TVar: " ^ (string_of_typ thy a));
-*)
-
-(* used by constr_list *)
-fun extract_constrs thy [] = [] |
-extract_constrs thy (a::r) =
-let
-fun delete_bold [] = []
-| delete_bold (x::xs) = if (is_prefix (op =) ("\^["::"["::"1"::"m"::[]) (x::xs))
- then (let val (_::_::_::s) = xs in delete_bold s end)
- else (if (is_prefix (op =) ("\^["::"["::"0"::"m"::[]) (x::xs))
- then (let val (_::_::_::s) = xs in delete_bold s end)
- else (x::delete_bold xs));
-fun delete_bold_string s = implode(delete_bold(explode s));
-(* from a constructor term in *.induct (with quantifiers,
-"Trueprop" and ?P stripped away) delivers the head *)
-fun extract_hd (_ $ Abs(_,_,r)) = extract_hd r |
-extract_hd (Const("Trueprop",_) $ r) = extract_hd r |
-extract_hd (Var(_,_) $ r) = extract_hd r |
-extract_hd (a $ b) = extract_hd a |
-extract_hd (Const(s,_)) = s |
-extract_hd _ = raise malformed;
-(* delivers constructor term string from a prem of *.induct *)
-fun extract_constr thy (_ $ Abs(a,T,r)) = extract_constr thy (snd(Syntax.variant_abs(a,T,r)))|
-extract_constr thy (Const("Trueprop",_) $ r) = extract_constr thy r |
-extract_constr thy (Var(_,_) $ r) = delete_bold_string(string_of_term thy r) |
-extract_constr _ _ = raise malformed;
-in
-(extract_hd a,extract_constr thy a) :: (extract_constrs thy r)
-end;
-
-(* delivering list of constructor terms of a datatype *)
-fun constr_list thy atyp =
-let
-fun act_name thy (Type(s,_)) = s |
-act_name _ s =
-error("malformed action type: " ^ (string_of_typ thy s));
-fun afpl ("." :: a) = [] |
-afpl [] = [] |
-afpl (a::r) = a :: (afpl r);
-fun unqualify s = implode(rev(afpl(rev(explode s))));
-val q_atypstr = act_name thy atyp;
-val uq_atypstr = unqualify q_atypstr;
-val prem = prems_of (PureThy.get_thm thy (uq_atypstr ^ ".induct"));
-in
-extract_constrs thy prem
-handle malformed =>
-error("malformed theorem : " ^ uq_atypstr ^ ".induct")
-end;
-
-fun check_for_constr thy atyp (a $ b) =
-let
-fun all_free (Free(_,_)) = true |
-all_free (a $ b) = if (all_free a) then (all_free b) else false |
-all_free _ = false;
-in
-if (all_free b) then (check_for_constr thy atyp a) else false
-end |
-check_for_constr thy atyp (Const(a,_)) =
-let
-val cl = constr_list thy atyp;
-fun fstmem a [] = false |
-fstmem a ((f,s)::r) = if (a=f) then true else (fstmem a r)
-in
-if (fstmem a cl) then true else false
-end |
-check_for_constr _ _ _ = false;
-
-(* delivering the free variables of a constructor term *)
-fun free_vars_of (t1 $ t2) = (free_vars_of t1) @ (free_vars_of t2) |
-free_vars_of (Const(_,_)) = [] |
-free_vars_of (Free(a,_)) = [a] |
-free_vars_of _ = raise malformed;
-
-(* making a constructor set from a constructor term (of signature) *)
-fun constr_set_string thy atyp ctstr =
-let
-val trm = OldGoals.simple_read_term thy atyp ctstr;
-val l = free_vars_of trm
-in
-if (check_for_constr thy atyp trm) then
-(if (l=[]) then ("{" ^ ctstr ^ "}")
-else "(UN " ^ (list_elements_of l) ^ ". {" ^ ctstr ^ "})")
-else (raise malformed)
-handle malformed =>
-error("malformed action term: " ^ (string_of_term thy trm))
-end;
-
-(* extracting constructor heads *)
-fun constructor_head thy atypstr s =
-let
-fun hd_of (Const(a,_)) = a |
-hd_of (t $ _) = hd_of t |
-hd_of _ = raise malformed;
-val trm = OldGoals.simple_read_term thy (Syntax.read_typ_global thy atypstr) s;
-in
-hd_of trm handle malformed =>
-error("malformed constructor of datatype " ^ atypstr ^ ": " ^ s)
-end;
-fun constructor_head_list _ _ [] = [] |
-constructor_head_list thy atypstr (a::r) =
- (constructor_head thy atypstr a)::(constructor_head_list thy atypstr r);
-
-(* producing an action set *)
-(*FIXME*)
-fun action_set_string thy atyp [] = "Set.empty" |
-action_set_string thy atyp [a] = constr_set_string thy atyp (strip_quote a) |
-action_set_string thy atyp (a::r) = (constr_set_string thy atyp (strip_quote a)) ^
- " Un " ^ (action_set_string thy atyp r);
-
-(* used by extend *)
-fun pstr s [] = "(" ^ s ^ "' = " ^ s ^ ")" |
-pstr s ((a,b)::r) =
-if (s=a) then ("(" ^ s ^ "' = (" ^ b ^ "))") else (pstr s r);
-fun poststring [] l = "" |
-poststring [(a,b)] l = pstr a l |
-poststring ((a,b)::r) l = (pstr a l) ^ " & " ^ (poststring r l);
-
-(* extends a (action string,condition,assignlist) tupel by a
-(action term,action string,condition,pseudo_condition,bool) tupel, used by extended_list
-(where bool indicates whether there is a precondition *)
-fun extend thy atyp statetupel (actstr,r,[]) =
-let
-val trm = OldGoals.simple_read_term thy atyp actstr;
-val rtrm = OldGoals.simple_read_term thy (Type("bool",[])) r;
-val flag = if (rtrm=Const("True",Type("bool",[]))) then false else true
-in
-if (check_for_constr thy atyp trm)
-then (trm,actstr, "(" ^ r ^ ") & " ^ (poststring statetupel []),r,flag)
-else
-error("transition " ^ actstr ^ " is not a pure constructor term")
-end |
-extend thy atyp statetupel (actstr,r,(a,b)::c) =
-let
-fun pseudo_poststring [] = "" |
-pseudo_poststring ((a,b)::[]) = "(" ^ a ^ " = (" ^ b ^ "))" |
-pseudo_poststring ((a,b)::r) = "(" ^ a ^ " = (" ^ b ^ ")) & " ^ (pseudo_poststring r);
-val trm = OldGoals.simple_read_term thy atyp actstr;
-val rtrm = OldGoals.simple_read_term thy (Type("bool",[])) r;
-val flag = if (rtrm=Const("True",Type("bool",[]))) then false else true
-in
-if (check_for_constr thy atyp trm) then
-(if ((a="") andalso (b="") andalso (c=[])) then (trm,actstr,r,"True",false)
-(* the case with transrel *)
- else
- (trm,actstr,"(" ^ r ^ ") & " ^ (poststring statetupel ((a,b)::c)),
- "(" ^ r ^ ") & " ^ (pseudo_poststring ((a,b)::c)),flag))
-else
-error("transition " ^ actstr ^ " is not a pure constructor term")
-end;
-(* used by make_alt_string *)
-fun extended_list _ _ _ [] = [] |
-extended_list thy atyp statetupel (a::r) =
- (extend thy atyp statetupel a) :: (extended_list thy atyp statetupel r);
-
-(* used by write_alts *)
-fun write_alt thy (chead,tr) inp out int [] =
-if (chead mem inp) then
-(
-error("Input action " ^ tr ^ " was not specified")
-) else (
-if (chead mem (out@int)) then
-(writeln("Action " ^ tr ^ " was completedly disabled due to missing specification")) else ();
-(tr ^ " => False",tr ^ " => False")) |
-write_alt thy (chead,ctrm) inp out int ((a,b,c,d,e)::r) =
-let
-fun hd_of (Const(a,_)) = a |
-hd_of (t $ _) = hd_of t |
-hd_of _ = raise malformed;
-fun occurs_again c [] = false |
-occurs_again c ((a,_,_,_,_)::r) = if (c=(hd_of a)) then true else (occurs_again c r);
-in
-if (chead=(hd_of a)) then
-(if ((chead mem inp) andalso e) then (
-error("Input action " ^ b ^ " has a precondition")
-) else (if (chead mem (inp@out@int)) then
- (if (occurs_again chead r) then (
-error("Two specifications for action: " ^ b)
- ) else (b ^ " => " ^ c,b ^ " => " ^ d))
- else (
-error("Action " ^ b ^ " is not in automaton signature")
-))) else (write_alt thy (chead,ctrm) inp out int r)
-handle malformed =>
-error ("malformed action term: " ^ (string_of_term thy a))
-end;
-
-(* used by make_alt_string *)
-fun write_alts thy (a,b) inp out int [] ttr = (a,b) |
-write_alts thy (a,b) inp out int [c] ttr =
-let
-val wa = write_alt thy c inp out int ttr
-in
- (a ^ (fst wa),b ^ (snd wa))
-end |
-write_alts thy (a,b) inp out int (c::r) ttr =
-let
-val wa = write_alt thy c inp out int ttr
-in
- write_alts thy (a ^ (fst wa) ^ " | ", b ^ (snd wa) ^ " | ") inp out int r ttr
-end;
-
-fun make_alt_string thy inp out int atyp statetupel trans =
-let
-val cl = constr_list thy atyp;
-val ttr = extended_list thy atyp statetupel trans;
-in
-write_alts thy ("","") inp out int cl ttr
-end;
-
-(* used in add_ioa *)
-fun check_free_primed (Free(a,_)) =
-let
-val (f::r) = rev(explode a)
-in
-if (f="'") then [a] else []
-end |
-check_free_primed (a $ b) = ((check_free_primed a) @ (check_free_primed b)) |
-check_free_primed (Abs(_,_,t)) = check_free_primed t |
-check_free_primed _ = [];
-
-fun overlap [] _ = true |
-overlap (a::r) l = if (a mem l) then (
-error("Two occurences of action " ^ a ^ " in automaton signature")
-) else (overlap r l);
-
-(* delivering some types of an automaton *)
-fun aut_type_of thy aut_name =
-let
-fun left_of (( _ $ left) $ _) = left |
-left_of _ = raise malformed;
-val aut_def = concl_of (PureThy.get_thm thy (aut_name ^ "_def"));
-in
-(#T(rep_cterm(cterm_of thy (left_of aut_def))))
-handle malformed => error ("malformed_theorem : " ^ aut_name ^ "_def")
-end;
-
-fun act_type_of thy (Type(pair1,(Type(pair_asig,(Type(actionset,[acttyp])::_))::_))) = acttyp |
-act_type_of thy t = error ("could not extract action type of following automaton type:\n" ^
-(string_of_typ thy t));
-fun st_type_of thy (Type(pair1,_::(Type(pair2,Type(initial_set,[statetyp])::_))::_)) = statetyp |
-st_type_of thy t = error ("could not extract state type of following automaton type:\n" ^
-(string_of_typ thy t));
-
-fun comp_st_type_of thy [a] = st_type_of thy (aut_type_of thy a) |
-comp_st_type_of thy (a::r) = Type("*",[st_type_of thy (aut_type_of thy a), comp_st_type_of thy r]) |
-comp_st_type_of _ _ = error "empty automaton list";
-
-(* checking consistency of action types (for composition) *)
-fun check_ac thy (a::r) =
-let
-fun ch_f_a thy acttyp [] = acttyp |
-ch_f_a thy acttyp (a::r) =
-let
-val auttyp = aut_type_of thy a;
-val ac = (act_type_of thy auttyp);
-in
-if (ac=acttyp) then (ch_f_a thy acttyp r) else (error "A")
-end;
-val auttyp = aut_type_of thy a;
-val acttyp = (act_type_of thy auttyp);
-in
-ch_f_a thy acttyp r
-end |
-check_ac _ [] = error "empty automaton list";
-
-fun clist [] = "" |
-clist [a] = a |
-clist (a::r) = a ^ " || " ^ (clist r);
-
-val add_defs = snd oo (PureThy.add_defs_cmd false o map Thm.no_attributes o map (apfst Binding.name));
-
-
-(* add_ioa *)
-
-fun add_ioa automaton_name action_type inp out int statetupel ini trans thy =
-(writeln("Constructing automaton " ^ automaton_name ^ " ...");
-let
-val state_type_string = type_product_of_varlist(statetupel);
-val styp = Syntax.read_typ_global thy state_type_string;
-val state_vars_tupel = "(" ^ (comma_list_of_varlist statetupel) ^ ")";
-val state_vars_primed = "(" ^ (primed_comma_list_of_varlist statetupel) ^ ")";
-val atyp = Syntax.read_typ_global thy action_type;
-val inp_set_string = action_set_string thy atyp inp;
-val out_set_string = action_set_string thy atyp out;
-val int_set_string = action_set_string thy atyp int;
-val inp_head_list = constructor_head_list thy action_type inp;
-val out_head_list = constructor_head_list thy action_type out;
-val int_head_list = constructor_head_list thy action_type int;
-val overlap_flag = ((overlap inp out) andalso (overlap inp int) andalso (overlap out int));
-val alt_string = make_alt_string thy inp_head_list out_head_list int_head_list
- atyp statetupel trans;
-val thy2 = (thy
-|> Sign.add_consts
-[(Binding.name (automaton_name ^ "_initial"), "(" ^ state_type_string ^ ")set" ,NoSyn),
-(Binding.name (automaton_name ^ "_asig"), "(" ^ action_type ^ ")signature" ,NoSyn),
-(Binding.name (automaton_name ^ "_trans"),
- "(" ^ action_type ^ "," ^ state_type_string ^ ")transition set" ,NoSyn),
-(Binding.name automaton_name, "(" ^ action_type ^ "," ^ state_type_string ^ ")ioa" ,NoSyn)]
-|> add_defs
-[(automaton_name ^ "_initial_def",
-automaton_name ^ "_initial == {" ^ state_vars_tupel ^ "." ^ ini ^ "}"),
-(automaton_name ^ "_asig_def",
-automaton_name ^ "_asig == (" ^
- inp_set_string ^ "," ^ out_set_string ^ "," ^ int_set_string ^ ")"),
-(automaton_name ^ "_trans_def",
-automaton_name ^ "_trans == {(" ^
- state_vars_tupel ^ ", act_of_" ^ automaton_name ^ ", " ^ state_vars_primed ^
-"). case act_of_" ^ automaton_name ^ " of " ^ fst(alt_string) ^ "}"),
-(automaton_name ^ "_def",
-automaton_name ^ " == (" ^ automaton_name ^ "_asig, " ^ automaton_name ^
-"_initial, " ^ automaton_name ^ "_trans,{},{})")
-])
-val chk_prime_list = (check_free_primed (OldGoals.simple_read_term thy2 (Type("bool",[]))
-( "case act_of_" ^ automaton_name ^ " of " ^ snd(alt_string))))
-in
-(
-if (chk_prime_list = []) then thy2
-else (
-error("Precondition or assignment terms in postconditions contain following primed variables:\n"
- ^ (list_elements_of chk_prime_list)))
-)
-end)
-
-fun add_composition automaton_name aut_list thy =
-(writeln("Constructing automaton " ^ automaton_name ^ " ...");
-let
-val acttyp = check_ac thy aut_list;
-val st_typ = comp_st_type_of thy aut_list;
-val comp_list = clist aut_list;
-in
-thy
-|> Sign.add_consts_i
-[(Binding.name automaton_name,
-Type("*",
-[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
- Type("*",[Type("set",[st_typ]),
- Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
- Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
-,NoSyn)]
-|> add_defs
-[(automaton_name ^ "_def",
-automaton_name ^ " == " ^ comp_list)]
-end)
-
-fun add_restriction automaton_name aut_source actlist thy =
-(writeln("Constructing automaton " ^ automaton_name ^ " ...");
-let
-val auttyp = aut_type_of thy aut_source;
-val acttyp = act_type_of thy auttyp;
-val rest_set = action_set_string thy acttyp actlist
-in
-thy
-|> Sign.add_consts_i
-[(Binding.name automaton_name, auttyp,NoSyn)]
-|> add_defs
-[(automaton_name ^ "_def",
-automaton_name ^ " == restrict " ^ aut_source ^ " " ^ rest_set)]
-end)
-fun add_hiding automaton_name aut_source actlist thy =
-(writeln("Constructing automaton " ^ automaton_name ^ " ...");
-let
-val auttyp = aut_type_of thy aut_source;
-val acttyp = act_type_of thy auttyp;
-val hid_set = action_set_string thy acttyp actlist
-in
-thy
-|> Sign.add_consts_i
-[(Binding.name automaton_name, auttyp,NoSyn)]
-|> add_defs
-[(automaton_name ^ "_def",
-automaton_name ^ " == hide " ^ aut_source ^ " " ^ hid_set)]
-end)
-
-fun ren_act_type_of thy funct =
- (case Term.fastype_of (Syntax.read_term_global thy funct) of
- Type ("fun", [a, b]) => a
- | _ => error "could not extract argument type of renaming function term");
-
-fun add_rename automaton_name aut_source fun_name thy =
-(writeln("Constructing automaton " ^ automaton_name ^ " ...");
-let
-val auttyp = aut_type_of thy aut_source;
-val st_typ = st_type_of thy auttyp;
-val acttyp = ren_act_type_of thy fun_name
-in
-thy
-|> Sign.add_consts_i
-[(Binding.name automaton_name,
-Type("*",
-[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
- Type("*",[Type("set",[st_typ]),
- Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
- Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
-,NoSyn)]
-|> add_defs
-[(automaton_name ^ "_def",
-automaton_name ^ " == rename " ^ aut_source ^ " (" ^ fun_name ^ ")")]
-end)
-
-
-(** outer syntax **)
-
-(* prepare results *)
-
-(*encoding transition specifications with a element of ParseTrans*)
-datatype ParseTrans = Rel of string | PP of string*(string*string)list;
-fun mk_trans_of_rel s = Rel(s);
-fun mk_trans_of_prepost (s,l) = PP(s,l);
-
-fun trans_of (a, Rel b) = (a, b, [("", "")])
- | trans_of (a, PP (b, l)) = (a, b, l);
-
-
-fun mk_ioa_decl (aut, ((((((action_type, inp), out), int), states), initial), trans)) =
- add_ioa aut action_type inp out int states initial (map trans_of trans);
-
-fun mk_composition_decl (aut, autlist) =
- add_composition aut autlist;
-
-fun mk_hiding_decl (aut, (actlist, source_aut)) =
- add_hiding aut source_aut actlist;
-
-fun mk_restriction_decl (aut, (source_aut, actlist)) =
- add_restriction aut source_aut actlist;
-
-fun mk_rename_decl (aut, (source_aut, rename_f)) =
- add_rename aut source_aut rename_f;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword ["signature", "actions", "inputs",
- "outputs", "internals", "states", "initially", "transitions", "pre",
- "post", "transrel", ":=", "compose", "hide_action", "in", "restrict", "to",
- "rename"];
-
-val actionlist = P.list1 P.term;
-val inputslist = P.$$$ "inputs" |-- P.!!! actionlist;
-val outputslist = P.$$$ "outputs" |-- P.!!! actionlist;
-val internalslist = P.$$$ "internals" |-- P.!!! actionlist;
-val stateslist = P.$$$ "states" |-- P.!!! (Scan.repeat1 (P.name --| P.$$$ "::" -- P.typ));
-val initial = P.$$$ "initially" |-- P.!!! P.term;
-val assign_list = P.list1 (P.name --| P.$$$ ":=" -- P.!!! P.term);
-val pre = P.$$$ "pre" |-- P.!!! P.term;
-val post = P.$$$ "post" |-- P.!!! assign_list;
-val pre1 = (pre -- (Scan.optional post [])) >> mk_trans_of_prepost;
-val post1 = ((Scan.optional pre "True") -- post) >> mk_trans_of_prepost;
-val transrel = (P.$$$ "transrel" |-- P.!!! P.term) >> mk_trans_of_rel;
-val transition = P.term -- (transrel || pre1 || post1);
-val translist = P.$$$ "transitions" |-- P.!!! (Scan.repeat1 transition);
-
-val ioa_decl =
- (P.name -- (P.$$$ "=" |--
- (P.$$$ "signature" |--
- P.!!! (P.$$$ "actions" |--
- (P.typ --
- (Scan.optional inputslist []) --
- (Scan.optional outputslist []) --
- (Scan.optional internalslist []) --
- stateslist --
- (Scan.optional initial "True") --
- translist))))) >> mk_ioa_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "compose" |-- P.!!! (P.list1 P.name))))
- >> mk_composition_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "hide_action" |--
- P.!!! (P.list1 P.term -- (P.$$$ "in" |-- P.name))))) >> mk_hiding_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "restrict" |--
- P.!!! (P.name -- (P.$$$ "to" |-- P.list1 P.term))))) >> mk_restriction_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "rename" |-- (P.!!! (P.name -- (P.$$$ "to" |-- P.term))))))
- >> mk_rename_decl;
-
-val _ =
- OuterSyntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" K.thy_decl
- (ioa_decl >> Toplevel.theory);
-
-end;
-
-end;
--- a/src/HOLCF/IsaMakefile Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/IsaMakefile Tue Jun 23 15:48:55 2009 +0100
@@ -67,8 +67,8 @@
Tools/domain/domain_library.ML \
Tools/domain/domain_syntax.ML \
Tools/domain/domain_theorems.ML \
- Tools/fixrec_package.ML \
- Tools/pcpodef_package.ML \
+ Tools/fixrec.ML \
+ Tools/pcpodef.ML \
holcf_logic.ML \
document/root.tex
@$(ISABELLE_TOOL) usedir -b -g true -r $(OUT)/HOL HOLCF
@@ -127,7 +127,7 @@
IOA/meta_theory/TL.thy IOA/meta_theory/TLS.thy \
IOA/meta_theory/LiveIOA.thy IOA/meta_theory/Pred.thy \
IOA/meta_theory/Abstraction.thy IOA/meta_theory/Simulations.thy \
- IOA/meta_theory/SimCorrectness.thy IOA/meta_theory/ioa_package.ML
+ IOA/meta_theory/SimCorrectness.thy IOA/meta_theory/automaton.ML
@cd IOA; $(ISABELLE_TOOL) usedir -b $(OUT)/HOLCF IOA
--- a/src/HOLCF/Pcpodef.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/Pcpodef.thy Tue Jun 23 15:48:55 2009 +0100
@@ -6,7 +6,7 @@
theory Pcpodef
imports Adm
-uses ("Tools/pcpodef_package.ML")
+uses ("Tools/pcpodef.ML")
begin
subsection {* Proving a subtype is a partial order *}
@@ -303,6 +303,6 @@
subsection {* HOLCF type definition package *}
-use "Tools/pcpodef_package.ML"
+use "Tools/pcpodef.ML"
end
--- a/src/HOLCF/Tools/domain/domain_axioms.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_axioms.ML Tue Jun 23 15:48:55 2009 +0100
@@ -6,7 +6,7 @@
signature DOMAIN_AXIOMS =
sig
- val copy_of_dtyp : (int -> term) -> DatatypeAux.dtyp -> term
+ val copy_of_dtyp : (int -> term) -> Datatype.dtyp -> term
val calc_axioms :
string -> Domain_Library.eq list -> int -> Domain_Library.eq ->
@@ -171,7 +171,7 @@
val mat_names = map mat_name con_names;
fun qualify n = Sign.full_name thy (Binding.name n);
val ms = map qualify con_names ~~ map qualify mat_names;
- in FixrecPackage.add_matchers ms thy end;
+ in Fixrec.add_matchers ms thy end;
fun add_axioms comp_dnam (eqs : eq list) thy' =
let
--- a/src/HOLCF/Tools/domain/domain_library.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_library.ML Tue Jun 23 15:48:55 2009 +0100
@@ -135,10 +135,10 @@
eqtype arg;
type cons = string * arg list;
type eq = (string * typ list) * cons list;
- val mk_arg : (bool * DatatypeAux.dtyp) * string option * string -> arg;
+ val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg;
val is_lazy : arg -> bool;
val rec_of : arg -> int;
- val dtyp_of : arg -> DatatypeAux.dtyp;
+ val dtyp_of : arg -> Datatype.dtyp;
val sel_of : arg -> string option;
val vname : arg -> string;
val upd_vname : (string -> string) -> arg -> arg;
@@ -154,7 +154,7 @@
val idx_name : 'a list -> string -> int -> string;
val app_rec_arg : (int -> term) -> arg -> term;
val con_app : string -> arg list -> term;
- val dtyp_of_eq : eq -> DatatypeAux.dtyp;
+ val dtyp_of_eq : eq -> Datatype.dtyp;
(* Name mangling *)
@@ -215,7 +215,7 @@
(* ----- constructor list handling ----- *)
type arg =
- (bool * DatatypeAux.dtyp) * (* (lazy, recursive element) *)
+ (bool * Datatype.dtyp) * (* (lazy, recursive element) *)
string option * (* selector name *)
string; (* argument name *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tools/fixrec.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,435 @@
+(* Title: HOLCF/Tools/fixrec.ML
+ Author: Amber Telfer and Brian Huffman
+
+Recursive function definition package for HOLCF.
+*)
+
+signature FIXREC =
+sig
+ val add_fixrec: bool -> (binding * typ option * mixfix) list
+ -> (Attrib.binding * term) list -> local_theory -> local_theory
+ val add_fixrec_cmd: bool -> (binding * string option * mixfix) list
+ -> (Attrib.binding * string) list -> local_theory -> local_theory
+ val add_fixpat: Thm.binding * term list -> theory -> theory
+ val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory
+ val add_matchers: (string * string) list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure Fixrec :> FIXREC =
+struct
+
+val def_cont_fix_eq = @{thm def_cont_fix_eq};
+val def_cont_fix_ind = @{thm def_cont_fix_ind};
+
+
+fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
+fun fixrec_eq_err thy s eq =
+ fixrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq));
+
+(*************************************************************************)
+(***************************** building types ****************************)
+(*************************************************************************)
+
+(* ->> is taken from holcf_logic.ML *)
+fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
+
+infixr 6 ->>; val (op ->>) = cfunT;
+
+fun cfunsT (Ts, U) = foldr cfunT U Ts;
+
+fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
+ | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
+
+fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U
+ | binder_cfun _ = [];
+
+fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U
+ | body_cfun T = T;
+
+fun strip_cfun T : typ list * typ =
+ (binder_cfun T, body_cfun T);
+
+fun maybeT T = Type(@{type_name "maybe"}, [T]);
+
+fun dest_maybeT (Type(@{type_name "maybe"}, [T])) = T
+ | dest_maybeT T = raise TYPE ("dest_maybeT", [T], []);
+
+fun tupleT [] = HOLogic.unitT
+ | tupleT [T] = T
+ | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts);
+
+fun matchT (T, U) =
+ body_cfun T ->> cfunsT (binder_cfun T, U) ->> U;
+
+
+(*************************************************************************)
+(***************************** building terms ****************************)
+(*************************************************************************)
+
+val mk_trp = HOLogic.mk_Trueprop;
+
+(* splits a cterm into the right and lefthand sides of equality *)
+fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t);
+
+(* similar to Thm.head_of, but for continuous application *)
+fun chead_of (Const(@{const_name Rep_CFun},_)$f$t) = chead_of f
+ | chead_of u = u;
+
+fun capply_const (S, T) =
+ Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T));
+
+fun cabs_const (S, T) =
+ Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T));
+
+fun mk_cabs t =
+ let val T = Term.fastype_of t
+ in cabs_const (Term.domain_type T, Term.range_type T) $ t end
+
+fun mk_capply (t, u) =
+ let val (S, T) =
+ case Term.fastype_of t of
+ Type(@{type_name "->"}, [S, T]) => (S, T)
+ | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]);
+ in capply_const (S, T) $ t $ u end;
+
+infix 0 ==; val (op ==) = Logic.mk_equals;
+infix 1 ===; val (op ===) = HOLogic.mk_eq;
+infix 9 ` ; val (op `) = mk_capply;
+
+(* builds the expression (LAM v. rhs) *)
+fun big_lambda v rhs =
+ cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs;
+
+(* builds the expression (LAM v1 v2 .. vn. rhs) *)
+fun big_lambdas [] rhs = rhs
+ | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
+
+fun mk_return t =
+ let val T = Term.fastype_of t
+ in Const(@{const_name Fixrec.return}, T ->> maybeT T) ` t end;
+
+fun mk_bind (t, u) =
+ let val (T, mU) = dest_cfunT (Term.fastype_of u);
+ val bindT = maybeT T ->> (T ->> mU) ->> mU;
+ in Const(@{const_name Fixrec.bind}, bindT) ` t ` u end;
+
+fun mk_mplus (t, u) =
+ let val mT = Term.fastype_of t
+ in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end;
+
+fun mk_run t =
+ let val mT = Term.fastype_of t
+ val T = dest_maybeT mT
+ in Const(@{const_name Fixrec.run}, mT ->> T) ` t end;
+
+fun mk_fix t =
+ let val (T, _) = dest_cfunT (Term.fastype_of t)
+ in Const(@{const_name fix}, (T ->> T) ->> T) ` t end;
+
+fun mk_cont t =
+ let val T = Term.fastype_of t
+ in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end;
+
+val mk_fst = HOLogic.mk_fst
+val mk_snd = HOLogic.mk_snd
+
+(* builds the expression (v1,v2,..,vn) *)
+fun mk_tuple [] = HOLogic.unit
+| mk_tuple (t::[]) = t
+| mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts);
+
+(* builds the expression (%(v1,v2,..,vn). rhs) *)
+fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
+ | lambda_tuple (v::[]) rhs = Term.lambda v rhs
+ | lambda_tuple (v::vs) rhs =
+ HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs));
+
+
+(*************************************************************************)
+(************* fixed-point definitions and unfolding theorems ************)
+(*************************************************************************)
+
+fun add_fixdefs
+ (fixes : ((binding * typ) * mixfix) list)
+ (spec : (Attrib.binding * term) list)
+ (lthy : local_theory) =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val names = map (Binding.name_of o fst o fst) fixes;
+ val all_names = space_implode "_" names;
+ val (lhss,rhss) = ListPair.unzip (map (dest_eqs o snd) spec);
+ val functional = lambda_tuple lhss (mk_tuple rhss);
+ val fixpoint = mk_fix (mk_cabs functional);
+
+ val cont_thm =
+ Goal.prove lthy [] [] (mk_trp (mk_cont functional))
+ (K (simp_tac (local_simpset_of lthy) 1));
+
+ fun one_def (l as Free(n,_)) r =
+ let val b = Long_Name.base_name n
+ in ((Binding.name (b^"_def"), []), r) end
+ | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form";
+ fun defs [] _ = []
+ | defs (l::[]) r = [one_def l r]
+ | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r);
+ val fixdefs = defs lhss fixpoint;
+ val define_all = fold_map (LocalTheory.define Thm.definitionK);
+ val (fixdef_thms : (term * (string * thm)) list, lthy') = lthy
+ |> define_all (map (apfst fst) fixes ~~ fixdefs);
+ fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2];
+ val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms);
+ val P = Var (("P", 0), map Term.fastype_of lhss ---> HOLogic.boolT);
+ val predicate = lambda_tuple lhss (list_comb (P, lhss));
+ val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm])
+ |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)]
+ |> LocalDefs.unfold lthy @{thms split_paired_all split_conv split_strict};
+ val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm])
+ |> LocalDefs.unfold lthy' @{thms split_conv};
+ fun unfolds [] thm = []
+ | unfolds (n::[]) thm = [(n^"_unfold", thm)]
+ | unfolds (n::ns) thm = let
+ val thmL = thm RS @{thm Pair_eqD1};
+ val thmR = thm RS @{thm Pair_eqD2};
+ in (n^"_unfold", thmL) :: unfolds ns thmR end;
+ val unfold_thms = unfolds names tuple_unfold_thm;
+ fun mk_note (n, thm) = ((Binding.name n, []), [thm]);
+ val (thmss, lthy'') = lthy'
+ |> fold_map (LocalTheory.note Thm.generatedK o mk_note)
+ ((all_names ^ "_induct", tuple_induct_thm) :: unfold_thms);
+ in
+ (lthy'', names, fixdef_thms, map snd unfold_thms)
+ end;
+
+(*************************************************************************)
+(*********** monadic notation and pattern matching compilation ***********)
+(*************************************************************************)
+
+structure FixrecMatchData = TheoryDataFun (
+ type T = string Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ tabs : T = Symtab.merge (K true) tabs;
+);
+
+(* associate match functions with pattern constants *)
+fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms);
+
+fun taken_names (t : term) : bstring list =
+ let
+ fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs
+ | taken (Free(a,_) , bs) = insert (op =) a bs
+ | taken (f $ u , bs) = taken (f, taken (u, bs))
+ | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs)
+ | taken (_ , bs) = bs;
+ in
+ taken (t, [])
+ end;
+
+(* builds a monadic term for matching a constructor pattern *)
+fun pre_build match_name pat rhs vs taken =
+ case pat of
+ Const(@{const_name Rep_CFun},_)$f$(v as Free(n,T)) =>
+ pre_build match_name f rhs (v::vs) taken
+ | Const(@{const_name Rep_CFun},_)$f$x =>
+ let val (rhs', v, taken') = pre_build match_name x rhs [] taken;
+ in pre_build match_name f rhs' (v::vs) taken' end
+ | Const(c,T) =>
+ let
+ val n = Name.variant taken "v";
+ fun result_type (Type(@{type_name "->"},[_,T])) (x::xs) = result_type T xs
+ | result_type T _ = T;
+ val v = Free(n, result_type T vs);
+ val m = Const(match_name c, matchT (T, fastype_of rhs));
+ val k = big_lambdas vs rhs;
+ in
+ (m`v`k, v, n::taken)
+ end
+ | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
+ | _ => fixrec_err "pre_build: invalid pattern";
+
+(* builds a monadic term for matching a function definition pattern *)
+(* returns (name, arity, matcher) *)
+fun building match_name pat rhs vs taken =
+ case pat of
+ Const(@{const_name Rep_CFun}, _)$f$(v as Free(n,T)) =>
+ building match_name f rhs (v::vs) taken
+ | Const(@{const_name Rep_CFun}, _)$f$x =>
+ let val (rhs', v, taken') = pre_build match_name x rhs [] taken;
+ in building match_name f rhs' (v::vs) taken' end
+ | Free(_,_) => ((pat, length vs), big_lambdas vs rhs)
+ | Const(_,_) => ((pat, length vs), big_lambdas vs rhs)
+ | _ => fixrec_err ("function is not declared as constant in theory: "
+ ^ ML_Syntax.print_term pat);
+
+fun strip_alls t =
+ if Logic.is_all t then strip_alls (snd (Logic.dest_all t)) else t;
+
+fun match_eq match_name eq =
+ let
+ val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq));
+ in
+ building match_name lhs (mk_return rhs) [] (taken_names eq)
+ end;
+
+(* returns the sum (using +++) of the terms in ms *)
+(* also applies "run" to the result! *)
+fun fatbar arity ms =
+ let
+ fun LAM_Ts 0 t = ([], Term.fastype_of t)
+ | LAM_Ts n (_ $ Abs(_,T,t)) =
+ let val (Ts, U) = LAM_Ts (n-1) t in (T::Ts, U) end
+ | LAM_Ts _ _ = fixrec_err "fatbar: internal error, not enough LAMs";
+ fun unLAM 0 t = t
+ | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t
+ | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs";
+ fun reLAM ([], U) t = t
+ | reLAM (T::Ts, U) t = reLAM (Ts, T ->> U) (cabs_const(T,U)$Abs("",T,t));
+ val msum = foldr1 mk_mplus (map (unLAM arity) ms);
+ val (Ts, U) = LAM_Ts arity (hd ms)
+ in
+ reLAM (rev Ts, dest_maybeT U) (mk_run msum)
+ end;
+
+(* this is the pattern-matching compiler function *)
+fun compile_pats match_name eqs =
+ let
+ val (((n::names),(a::arities)),mats) =
+ apfst ListPair.unzip (ListPair.unzip (map (match_eq match_name) eqs));
+ val cname = if forall (fn x => n=x) names then n
+ else fixrec_err "all equations in block must define the same function";
+ val arity = if forall (fn x => a=x) arities then a
+ else fixrec_err "all equations in block must have the same arity";
+ val rhs = fatbar arity mats;
+ in
+ mk_trp (cname === rhs)
+ end;
+
+(*************************************************************************)
+(********************** Proving associated theorems **********************)
+(*************************************************************************)
+
+(* proves a block of pattern matching equations as theorems, using unfold *)
+fun make_simps lthy (unfold_thm, eqns : (Attrib.binding * term) list) =
+ let
+ val tacs =
+ [rtac (unfold_thm RS @{thm ssubst_lhs}) 1,
+ asm_simp_tac (local_simpset_of lthy) 1];
+ fun prove_term t = Goal.prove lthy [] [] t (K (EVERY tacs));
+ fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t);
+ in
+ map prove_eqn eqns
+ end;
+
+(*************************************************************************)
+(************************* Main fixrec function **************************)
+(*************************************************************************)
+
+local
+(* code adapted from HOL/Tools/primrec.ML *)
+
+fun gen_fixrec
+ (set_group : bool)
+ prep_spec
+ (strict : bool)
+ raw_fixes
+ raw_spec
+ (lthy : local_theory) =
+ let
+ val (fixes : ((binding * typ) * mixfix) list,
+ spec : (Attrib.binding * term) list) =
+ fst (prep_spec raw_fixes raw_spec lthy);
+ val chead_of_spec =
+ chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd;
+ fun name_of (Free (n, _)) = n
+ | name_of t = fixrec_err ("unknown term");
+ val all_names = map (name_of o chead_of_spec) spec;
+ val names = distinct (op =) all_names;
+ fun block_of_name n =
+ map_filter
+ (fn (m,eq) => if m = n then SOME eq else NONE)
+ (all_names ~~ spec);
+ val blocks = map block_of_name names;
+
+ val matcher_tab = FixrecMatchData.get (ProofContext.theory_of lthy);
+ fun match_name c =
+ case Symtab.lookup matcher_tab c of SOME m => m
+ | NONE => fixrec_err ("unknown pattern constructor: " ^ c);
+
+ val matches = map (compile_pats match_name) (map (map snd) blocks);
+ val spec' = map (pair Attrib.empty_binding) matches;
+ val (lthy', cnames, fixdef_thms, unfold_thms) =
+ add_fixdefs fixes spec' lthy;
+ in
+ if strict then let (* only prove simp rules if strict = true *)
+ val simps : (Attrib.binding * thm) list list =
+ map (make_simps lthy') (unfold_thms ~~ blocks);
+ fun mk_bind n : Attrib.binding =
+ (Binding.name (n ^ "_simps"),
+ [Attrib.internal (K Simplifier.simp_add)]);
+ val simps1 : (Attrib.binding * thm list) list =
+ map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps);
+ val simps2 : (Attrib.binding * thm list) list =
+ map (apsnd (fn thm => [thm])) (List.concat simps);
+ val (_, lthy'') = lthy'
+ |> fold_map (LocalTheory.note Thm.generatedK) (simps1 @ simps2);
+ in
+ lthy''
+ end
+ else lthy'
+ end;
+
+in
+
+val add_fixrec = gen_fixrec false Specification.check_spec;
+val add_fixrec_cmd = gen_fixrec true Specification.read_spec;
+
+end; (* local *)
+
+(*************************************************************************)
+(******************************** Fixpat *********************************)
+(*************************************************************************)
+
+fun fix_pat thy t =
+ let
+ val T = fastype_of t;
+ val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
+ val cname = case chead_of t of Const(c,_) => c | _ =>
+ fixrec_err "function is not declared as constant in theory";
+ val unfold_thm = PureThy.get_thm thy (cname^"_unfold");
+ val simp = Goal.prove_global thy [] [] eq
+ (fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]);
+ in simp end;
+
+fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy =
+ let
+ val atts = map (prep_attrib thy) srcs;
+ val ts = map (prep_term thy) strings;
+ val simps = map (fix_pat thy) ts;
+ in
+ (snd o PureThy.add_thmss [((name, simps), atts)]) thy
+ end;
+
+val add_fixpat = gen_add_fixpat Sign.cert_term (K I);
+val add_fixpat_cmd = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
+
+
+(*************************************************************************)
+(******************************** Parsers ********************************)
+(*************************************************************************)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl
+ ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs
+ >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs));
+
+val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
+ (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd));
+
+end;
+
+val setup = FixrecMatchData.init;
+
+end;
--- a/src/HOLCF/Tools/fixrec_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,435 +0,0 @@
-(* Title: HOLCF/Tools/fixrec_package.ML
- Author: Amber Telfer and Brian Huffman
-
-Recursive function definition package for HOLCF.
-*)
-
-signature FIXREC_PACKAGE =
-sig
- val add_fixrec: bool -> (binding * typ option * mixfix) list
- -> (Attrib.binding * term) list -> local_theory -> local_theory
- val add_fixrec_cmd: bool -> (binding * string option * mixfix) list
- -> (Attrib.binding * string) list -> local_theory -> local_theory
- val add_fixpat: Thm.binding * term list -> theory -> theory
- val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory
- val add_matchers: (string * string) list -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure FixrecPackage :> FIXREC_PACKAGE =
-struct
-
-val def_cont_fix_eq = @{thm def_cont_fix_eq};
-val def_cont_fix_ind = @{thm def_cont_fix_ind};
-
-
-fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
-fun fixrec_eq_err thy s eq =
- fixrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq));
-
-(*************************************************************************)
-(***************************** building types ****************************)
-(*************************************************************************)
-
-(* ->> is taken from holcf_logic.ML *)
-fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
-
-infixr 6 ->>; val (op ->>) = cfunT;
-
-fun cfunsT (Ts, U) = foldr cfunT U Ts;
-
-fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
- | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
-
-fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U
- | binder_cfun _ = [];
-
-fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U
- | body_cfun T = T;
-
-fun strip_cfun T : typ list * typ =
- (binder_cfun T, body_cfun T);
-
-fun maybeT T = Type(@{type_name "maybe"}, [T]);
-
-fun dest_maybeT (Type(@{type_name "maybe"}, [T])) = T
- | dest_maybeT T = raise TYPE ("dest_maybeT", [T], []);
-
-fun tupleT [] = HOLogic.unitT
- | tupleT [T] = T
- | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts);
-
-fun matchT (T, U) =
- body_cfun T ->> cfunsT (binder_cfun T, U) ->> U;
-
-
-(*************************************************************************)
-(***************************** building terms ****************************)
-(*************************************************************************)
-
-val mk_trp = HOLogic.mk_Trueprop;
-
-(* splits a cterm into the right and lefthand sides of equality *)
-fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t);
-
-(* similar to Thm.head_of, but for continuous application *)
-fun chead_of (Const(@{const_name Rep_CFun},_)$f$t) = chead_of f
- | chead_of u = u;
-
-fun capply_const (S, T) =
- Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T));
-
-fun cabs_const (S, T) =
- Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T));
-
-fun mk_cabs t =
- let val T = Term.fastype_of t
- in cabs_const (Term.domain_type T, Term.range_type T) $ t end
-
-fun mk_capply (t, u) =
- let val (S, T) =
- case Term.fastype_of t of
- Type(@{type_name "->"}, [S, T]) => (S, T)
- | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]);
- in capply_const (S, T) $ t $ u end;
-
-infix 0 ==; val (op ==) = Logic.mk_equals;
-infix 1 ===; val (op ===) = HOLogic.mk_eq;
-infix 9 ` ; val (op `) = mk_capply;
-
-(* builds the expression (LAM v. rhs) *)
-fun big_lambda v rhs =
- cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs;
-
-(* builds the expression (LAM v1 v2 .. vn. rhs) *)
-fun big_lambdas [] rhs = rhs
- | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
-
-fun mk_return t =
- let val T = Term.fastype_of t
- in Const(@{const_name Fixrec.return}, T ->> maybeT T) ` t end;
-
-fun mk_bind (t, u) =
- let val (T, mU) = dest_cfunT (Term.fastype_of u);
- val bindT = maybeT T ->> (T ->> mU) ->> mU;
- in Const(@{const_name Fixrec.bind}, bindT) ` t ` u end;
-
-fun mk_mplus (t, u) =
- let val mT = Term.fastype_of t
- in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end;
-
-fun mk_run t =
- let val mT = Term.fastype_of t
- val T = dest_maybeT mT
- in Const(@{const_name Fixrec.run}, mT ->> T) ` t end;
-
-fun mk_fix t =
- let val (T, _) = dest_cfunT (Term.fastype_of t)
- in Const(@{const_name fix}, (T ->> T) ->> T) ` t end;
-
-fun mk_cont t =
- let val T = Term.fastype_of t
- in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end;
-
-val mk_fst = HOLogic.mk_fst
-val mk_snd = HOLogic.mk_snd
-
-(* builds the expression (v1,v2,..,vn) *)
-fun mk_tuple [] = HOLogic.unit
-| mk_tuple (t::[]) = t
-| mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts);
-
-(* builds the expression (%(v1,v2,..,vn). rhs) *)
-fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
- | lambda_tuple (v::[]) rhs = Term.lambda v rhs
- | lambda_tuple (v::vs) rhs =
- HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs));
-
-
-(*************************************************************************)
-(************* fixed-point definitions and unfolding theorems ************)
-(*************************************************************************)
-
-fun add_fixdefs
- (fixes : ((binding * typ) * mixfix) list)
- (spec : (Attrib.binding * term) list)
- (lthy : local_theory) =
- let
- val thy = ProofContext.theory_of lthy;
- val names = map (Binding.name_of o fst o fst) fixes;
- val all_names = space_implode "_" names;
- val (lhss,rhss) = ListPair.unzip (map (dest_eqs o snd) spec);
- val functional = lambda_tuple lhss (mk_tuple rhss);
- val fixpoint = mk_fix (mk_cabs functional);
-
- val cont_thm =
- Goal.prove lthy [] [] (mk_trp (mk_cont functional))
- (K (simp_tac (local_simpset_of lthy) 1));
-
- fun one_def (l as Free(n,_)) r =
- let val b = Long_Name.base_name n
- in ((Binding.name (b^"_def"), []), r) end
- | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form";
- fun defs [] _ = []
- | defs (l::[]) r = [one_def l r]
- | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r);
- val fixdefs = defs lhss fixpoint;
- val define_all = fold_map (LocalTheory.define Thm.definitionK);
- val (fixdef_thms : (term * (string * thm)) list, lthy') = lthy
- |> define_all (map (apfst fst) fixes ~~ fixdefs);
- fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2];
- val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms);
- val P = Var (("P", 0), map Term.fastype_of lhss ---> HOLogic.boolT);
- val predicate = lambda_tuple lhss (list_comb (P, lhss));
- val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm])
- |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)]
- |> LocalDefs.unfold lthy @{thms split_paired_all split_conv split_strict};
- val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm])
- |> LocalDefs.unfold lthy' @{thms split_conv};
- fun unfolds [] thm = []
- | unfolds (n::[]) thm = [(n^"_unfold", thm)]
- | unfolds (n::ns) thm = let
- val thmL = thm RS @{thm Pair_eqD1};
- val thmR = thm RS @{thm Pair_eqD2};
- in (n^"_unfold", thmL) :: unfolds ns thmR end;
- val unfold_thms = unfolds names tuple_unfold_thm;
- fun mk_note (n, thm) = ((Binding.name n, []), [thm]);
- val (thmss, lthy'') = lthy'
- |> fold_map (LocalTheory.note Thm.generatedK o mk_note)
- ((all_names ^ "_induct", tuple_induct_thm) :: unfold_thms);
- in
- (lthy'', names, fixdef_thms, map snd unfold_thms)
- end;
-
-(*************************************************************************)
-(*********** monadic notation and pattern matching compilation ***********)
-(*************************************************************************)
-
-structure FixrecMatchData = TheoryDataFun (
- type T = string Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ tabs : T = Symtab.merge (K true) tabs;
-);
-
-(* associate match functions with pattern constants *)
-fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms);
-
-fun taken_names (t : term) : bstring list =
- let
- fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs
- | taken (Free(a,_) , bs) = insert (op =) a bs
- | taken (f $ u , bs) = taken (f, taken (u, bs))
- | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs)
- | taken (_ , bs) = bs;
- in
- taken (t, [])
- end;
-
-(* builds a monadic term for matching a constructor pattern *)
-fun pre_build match_name pat rhs vs taken =
- case pat of
- Const(@{const_name Rep_CFun},_)$f$(v as Free(n,T)) =>
- pre_build match_name f rhs (v::vs) taken
- | Const(@{const_name Rep_CFun},_)$f$x =>
- let val (rhs', v, taken') = pre_build match_name x rhs [] taken;
- in pre_build match_name f rhs' (v::vs) taken' end
- | Const(c,T) =>
- let
- val n = Name.variant taken "v";
- fun result_type (Type(@{type_name "->"},[_,T])) (x::xs) = result_type T xs
- | result_type T _ = T;
- val v = Free(n, result_type T vs);
- val m = Const(match_name c, matchT (T, fastype_of rhs));
- val k = big_lambdas vs rhs;
- in
- (m`v`k, v, n::taken)
- end
- | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
- | _ => fixrec_err "pre_build: invalid pattern";
-
-(* builds a monadic term for matching a function definition pattern *)
-(* returns (name, arity, matcher) *)
-fun building match_name pat rhs vs taken =
- case pat of
- Const(@{const_name Rep_CFun}, _)$f$(v as Free(n,T)) =>
- building match_name f rhs (v::vs) taken
- | Const(@{const_name Rep_CFun}, _)$f$x =>
- let val (rhs', v, taken') = pre_build match_name x rhs [] taken;
- in building match_name f rhs' (v::vs) taken' end
- | Free(_,_) => ((pat, length vs), big_lambdas vs rhs)
- | Const(_,_) => ((pat, length vs), big_lambdas vs rhs)
- | _ => fixrec_err ("function is not declared as constant in theory: "
- ^ ML_Syntax.print_term pat);
-
-fun strip_alls t =
- if Logic.is_all t then strip_alls (snd (Logic.dest_all t)) else t;
-
-fun match_eq match_name eq =
- let
- val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq));
- in
- building match_name lhs (mk_return rhs) [] (taken_names eq)
- end;
-
-(* returns the sum (using +++) of the terms in ms *)
-(* also applies "run" to the result! *)
-fun fatbar arity ms =
- let
- fun LAM_Ts 0 t = ([], Term.fastype_of t)
- | LAM_Ts n (_ $ Abs(_,T,t)) =
- let val (Ts, U) = LAM_Ts (n-1) t in (T::Ts, U) end
- | LAM_Ts _ _ = fixrec_err "fatbar: internal error, not enough LAMs";
- fun unLAM 0 t = t
- | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t
- | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs";
- fun reLAM ([], U) t = t
- | reLAM (T::Ts, U) t = reLAM (Ts, T ->> U) (cabs_const(T,U)$Abs("",T,t));
- val msum = foldr1 mk_mplus (map (unLAM arity) ms);
- val (Ts, U) = LAM_Ts arity (hd ms)
- in
- reLAM (rev Ts, dest_maybeT U) (mk_run msum)
- end;
-
-(* this is the pattern-matching compiler function *)
-fun compile_pats match_name eqs =
- let
- val (((n::names),(a::arities)),mats) =
- apfst ListPair.unzip (ListPair.unzip (map (match_eq match_name) eqs));
- val cname = if forall (fn x => n=x) names then n
- else fixrec_err "all equations in block must define the same function";
- val arity = if forall (fn x => a=x) arities then a
- else fixrec_err "all equations in block must have the same arity";
- val rhs = fatbar arity mats;
- in
- mk_trp (cname === rhs)
- end;
-
-(*************************************************************************)
-(********************** Proving associated theorems **********************)
-(*************************************************************************)
-
-(* proves a block of pattern matching equations as theorems, using unfold *)
-fun make_simps lthy (unfold_thm, eqns : (Attrib.binding * term) list) =
- let
- val tacs =
- [rtac (unfold_thm RS @{thm ssubst_lhs}) 1,
- asm_simp_tac (local_simpset_of lthy) 1];
- fun prove_term t = Goal.prove lthy [] [] t (K (EVERY tacs));
- fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t);
- in
- map prove_eqn eqns
- end;
-
-(*************************************************************************)
-(************************* Main fixrec function **************************)
-(*************************************************************************)
-
-local
-(* code adapted from HOL/Tools/primrec_package.ML *)
-
-fun gen_fixrec
- (set_group : bool)
- prep_spec
- (strict : bool)
- raw_fixes
- raw_spec
- (lthy : local_theory) =
- let
- val (fixes : ((binding * typ) * mixfix) list,
- spec : (Attrib.binding * term) list) =
- fst (prep_spec raw_fixes raw_spec lthy);
- val chead_of_spec =
- chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd;
- fun name_of (Free (n, _)) = n
- | name_of t = fixrec_err ("unknown term");
- val all_names = map (name_of o chead_of_spec) spec;
- val names = distinct (op =) all_names;
- fun block_of_name n =
- map_filter
- (fn (m,eq) => if m = n then SOME eq else NONE)
- (all_names ~~ spec);
- val blocks = map block_of_name names;
-
- val matcher_tab = FixrecMatchData.get (ProofContext.theory_of lthy);
- fun match_name c =
- case Symtab.lookup matcher_tab c of SOME m => m
- | NONE => fixrec_err ("unknown pattern constructor: " ^ c);
-
- val matches = map (compile_pats match_name) (map (map snd) blocks);
- val spec' = map (pair Attrib.empty_binding) matches;
- val (lthy', cnames, fixdef_thms, unfold_thms) =
- add_fixdefs fixes spec' lthy;
- in
- if strict then let (* only prove simp rules if strict = true *)
- val simps : (Attrib.binding * thm) list list =
- map (make_simps lthy') (unfold_thms ~~ blocks);
- fun mk_bind n : Attrib.binding =
- (Binding.name (n ^ "_simps"),
- [Attrib.internal (K Simplifier.simp_add)]);
- val simps1 : (Attrib.binding * thm list) list =
- map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps);
- val simps2 : (Attrib.binding * thm list) list =
- map (apsnd (fn thm => [thm])) (List.concat simps);
- val (_, lthy'') = lthy'
- |> fold_map (LocalTheory.note Thm.generatedK) (simps1 @ simps2);
- in
- lthy''
- end
- else lthy'
- end;
-
-in
-
-val add_fixrec = gen_fixrec false Specification.check_spec;
-val add_fixrec_cmd = gen_fixrec true Specification.read_spec;
-
-end; (* local *)
-
-(*************************************************************************)
-(******************************** Fixpat *********************************)
-(*************************************************************************)
-
-fun fix_pat thy t =
- let
- val T = fastype_of t;
- val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
- val cname = case chead_of t of Const(c,_) => c | _ =>
- fixrec_err "function is not declared as constant in theory";
- val unfold_thm = PureThy.get_thm thy (cname^"_unfold");
- val simp = Goal.prove_global thy [] [] eq
- (fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]);
- in simp end;
-
-fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy =
- let
- val atts = map (prep_attrib thy) srcs;
- val ts = map (prep_term thy) strings;
- val simps = map (fix_pat thy) ts;
- in
- (snd o PureThy.add_thmss [((name, simps), atts)]) thy
- end;
-
-val add_fixpat = gen_add_fixpat Sign.cert_term (K I);
-val add_fixpat_cmd = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
-
-
-(*************************************************************************)
-(******************************** Parsers ********************************)
-(*************************************************************************)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl
- ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs
- >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs));
-
-val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
- (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd));
-
-end;
-
-val setup = FixrecMatchData.init;
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tools/pcpodef.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,201 @@
+(* Title: HOLCF/Tools/pcpodef.ML
+ Author: Brian Huffman
+
+Primitive domain definitions for HOLCF, similar to Gordon/HOL-style
+typedef (see also ~~/src/HOL/Tools/typedef.ML).
+*)
+
+signature PCPODEF =
+sig
+ val pcpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
+ * (binding * binding) option -> theory -> Proof.state
+ val pcpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
+ * (binding * binding) option -> theory -> Proof.state
+ val cpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
+ * (binding * binding) option -> theory -> Proof.state
+ val cpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
+ * (binding * binding) option -> theory -> Proof.state
+end;
+
+structure Pcpodef :> PCPODEF =
+struct
+
+(** type definitions **)
+
+(* prepare_cpodef *)
+
+fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
+
+fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT);
+fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P);
+
+fun prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy =
+ let
+ val _ = Theory.requires thy "Pcpodef" "pcpodefs";
+ val ctxt = ProofContext.init thy;
+
+ val full = Sign.full_name thy;
+ val full_name = full name;
+ val bname = Binding.name_of name;
+
+ (*rhs*)
+ val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
+ val setT = Term.fastype_of set;
+ val rhs_tfrees = Term.add_tfrees set [];
+ val oldT = HOLogic.dest_setT setT handle TYPE _ =>
+ error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
+
+ (*goal*)
+ val goal_UU_mem = HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set));
+ val goal_nonempty =
+ HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
+ val goal_admissible =
+ HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
+
+ (*lhs*)
+ val defS = Sign.defaultS thy;
+ val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
+ val lhs_sorts = map snd lhs_tfrees;
+
+ val tname = Binding.map_name (Syntax.type_name mx) t;
+ val full_tname = full tname;
+ val newT = Type (full_tname, map TFree lhs_tfrees);
+
+ val (Rep_name, Abs_name) =
+ (case opt_morphs of
+ NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
+ | SOME morphs => morphs);
+ val RepC = Const (full Rep_name, newT --> oldT);
+ fun belowC T = Const (@{const_name below}, T --> T --> HOLogic.boolT);
+ val below_def = Logic.mk_equals (belowC newT,
+ Abs ("x", newT, Abs ("y", newT, belowC oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))));
+
+ fun make_po tac thy1 =
+ let
+ val ((_, {type_definition, set_def, ...}), thy2) = thy1
+ |> Typedef.add_typedef def (SOME name) (t, vs, mx) set opt_morphs tac;
+ val lthy3 = thy2
+ |> TheoryTarget.instantiation ([full_tname], lhs_tfrees, @{sort po});
+ val below_def' = Syntax.check_term lthy3 below_def;
+ val ((_, (_, below_definition')), lthy4) = lthy3
+ |> Specification.definition (NONE,
+ ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_def'));
+ val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4);
+ val below_definition = singleton (ProofContext.export lthy4 ctxt_thy) below_definition';
+ val thy5 = lthy4
+ |> Class.prove_instantiation_instance
+ (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_definition]) 1))
+ |> LocalTheory.exit_global;
+ in ((type_definition, below_definition, set_def), thy5) end;
+
+ fun make_cpo admissible (type_def, below_def, set_def) theory =
+ let
+ val admissible' = fold_rule (the_list set_def) admissible;
+ val cpo_thms = map (Thm.transfer theory) [type_def, below_def, admissible'];
+ val theory' = theory
+ |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo})
+ (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
+ val cpo_thms' = map (Thm.transfer theory') cpo_thms;
+ in
+ theory'
+ |> Sign.add_path (Binding.name_of name)
+ |> PureThy.add_thms
+ ([((Binding.prefix_name "adm_" name, admissible'), []),
+ ((Binding.prefix_name "cont_" Rep_name, @{thm typedef_cont_Rep} OF cpo_thms'), []),
+ ((Binding.prefix_name "cont_" Abs_name, @{thm typedef_cont_Abs} OF cpo_thms'), []),
+ ((Binding.prefix_name "lub_" name, @{thm typedef_lub} OF cpo_thms'), []),
+ ((Binding.prefix_name "thelub_" name, @{thm typedef_thelub} OF cpo_thms'), []),
+ ((Binding.prefix_name "compact_" name, @{thm typedef_compact} OF cpo_thms'), [])])
+ |> snd
+ |> Sign.parent_path
+ end;
+
+ fun make_pcpo UU_mem (type_def, below_def, set_def) theory =
+ let
+ val UU_mem' = fold_rule (the_list set_def) UU_mem;
+ val pcpo_thms = map (Thm.transfer theory) [type_def, below_def, UU_mem'];
+ val theory' = theory
+ |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
+ (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
+ val pcpo_thms' = map (Thm.transfer theory') pcpo_thms;
+ in
+ theory'
+ |> Sign.add_path (Binding.name_of name)
+ |> PureThy.add_thms
+ ([((Binding.suffix_name "_strict" Rep_name, @{thm typedef_Rep_strict} OF pcpo_thms'), []),
+ ((Binding.suffix_name "_strict" Abs_name, @{thm typedef_Abs_strict} OF pcpo_thms'), []),
+ ((Binding.suffix_name "_strict_iff" Rep_name, @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
+ ((Binding.suffix_name "_strict_iff" Abs_name, @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
+ ((Binding.suffix_name "_defined" Rep_name, @{thm typedef_Rep_defined} OF pcpo_thms'), []),
+ ((Binding.suffix_name "_defined" Abs_name, @{thm typedef_Abs_defined} OF pcpo_thms'), [])])
+ |> snd
+ |> Sign.parent_path
+ end;
+
+ fun pcpodef_result UU_mem admissible =
+ make_po (Tactic.rtac exI 1 THEN Tactic.rtac UU_mem 1)
+ #-> (fn defs => make_cpo admissible defs #> make_pcpo UU_mem defs);
+
+ fun cpodef_result nonempty admissible =
+ make_po (Tactic.rtac nonempty 1)
+ #-> make_cpo admissible;
+ in
+ if pcpo
+ then (goal_UU_mem, goal_admissible, pcpodef_result)
+ else (goal_nonempty, goal_admissible, cpodef_result)
+ end
+ handle ERROR msg =>
+ cat_error msg ("The error(s) above occurred in cpodef " ^ quote (Binding.str_of name));
+
+
+(* proof interface *)
+
+local
+
+fun gen_pcpodef_proof prep_term pcpo ((def, name), typ, set, opt_morphs) thy =
+ let
+ val (goal1, goal2, make_result) =
+ prepare_pcpodef prep_term pcpo def name typ set opt_morphs thy;
+ fun after_qed [[th1, th2]] = ProofContext.theory (make_result th1 th2);
+ in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] (ProofContext.init thy) end;
+
+in
+
+fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term true x;
+fun pcpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term true x;
+
+fun cpodef_proof x = gen_pcpodef_proof Syntax.check_term false x;
+fun cpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term false x;
+
+end;
+
+
+
+(** outer syntax **)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val typedef_proof_decl =
+ Scan.optional (P.$$$ "(" |--
+ ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
+ --| P.$$$ ")") (true, NONE) --
+ (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
+ Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
+
+fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
+ (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
+ ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs);
+
+val _ =
+ OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
+ (typedef_proof_decl >>
+ (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true)));
+
+val _ =
+ OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
+ (typedef_proof_decl >>
+ (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false)));
+
+end;
+
+end;
--- a/src/HOLCF/Tools/pcpodef_package.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,201 +0,0 @@
-(* Title: HOLCF/Tools/pcpodef_package.ML
- Author: Brian Huffman
-
-Primitive domain definitions for HOLCF, similar to Gordon/HOL-style
-typedef (see also ~~/src/HOL/Tools/typedef_package.ML).
-*)
-
-signature PCPODEF_PACKAGE =
-sig
- val pcpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
- * (binding * binding) option -> theory -> Proof.state
- val pcpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
- * (binding * binding) option -> theory -> Proof.state
- val cpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
- * (binding * binding) option -> theory -> Proof.state
- val cpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
- * (binding * binding) option -> theory -> Proof.state
-end;
-
-structure PcpodefPackage :> PCPODEF_PACKAGE =
-struct
-
-(** type definitions **)
-
-(* prepare_cpodef *)
-
-fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
-
-fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT);
-fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P);
-
-fun prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy =
- let
- val _ = Theory.requires thy "Pcpodef" "pcpodefs";
- val ctxt = ProofContext.init thy;
-
- val full = Sign.full_name thy;
- val full_name = full name;
- val bname = Binding.name_of name;
-
- (*rhs*)
- val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
- val setT = Term.fastype_of set;
- val rhs_tfrees = Term.add_tfrees set [];
- val oldT = HOLogic.dest_setT setT handle TYPE _ =>
- error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
-
- (*goal*)
- val goal_UU_mem = HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set));
- val goal_nonempty =
- HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
- val goal_admissible =
- HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
-
- (*lhs*)
- val defS = Sign.defaultS thy;
- val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
- val lhs_sorts = map snd lhs_tfrees;
-
- val tname = Binding.map_name (Syntax.type_name mx) t;
- val full_tname = full tname;
- val newT = Type (full_tname, map TFree lhs_tfrees);
-
- val (Rep_name, Abs_name) =
- (case opt_morphs of
- NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
- | SOME morphs => morphs);
- val RepC = Const (full Rep_name, newT --> oldT);
- fun belowC T = Const (@{const_name below}, T --> T --> HOLogic.boolT);
- val below_def = Logic.mk_equals (belowC newT,
- Abs ("x", newT, Abs ("y", newT, belowC oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))));
-
- fun make_po tac thy1 =
- let
- val ((_, {type_definition, set_def, ...}), thy2) = thy1
- |> TypedefPackage.add_typedef def (SOME name) (t, vs, mx) set opt_morphs tac;
- val lthy3 = thy2
- |> TheoryTarget.instantiation ([full_tname], lhs_tfrees, @{sort po});
- val below_def' = Syntax.check_term lthy3 below_def;
- val ((_, (_, below_definition')), lthy4) = lthy3
- |> Specification.definition (NONE,
- ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_def'));
- val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4);
- val below_definition = singleton (ProofContext.export lthy4 ctxt_thy) below_definition';
- val thy5 = lthy4
- |> Class.prove_instantiation_instance
- (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_definition]) 1))
- |> LocalTheory.exit_global;
- in ((type_definition, below_definition, set_def), thy5) end;
-
- fun make_cpo admissible (type_def, below_def, set_def) theory =
- let
- val admissible' = fold_rule (the_list set_def) admissible;
- val cpo_thms = map (Thm.transfer theory) [type_def, below_def, admissible'];
- val theory' = theory
- |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo})
- (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
- val cpo_thms' = map (Thm.transfer theory') cpo_thms;
- in
- theory'
- |> Sign.add_path (Binding.name_of name)
- |> PureThy.add_thms
- ([((Binding.prefix_name "adm_" name, admissible'), []),
- ((Binding.prefix_name "cont_" Rep_name, @{thm typedef_cont_Rep} OF cpo_thms'), []),
- ((Binding.prefix_name "cont_" Abs_name, @{thm typedef_cont_Abs} OF cpo_thms'), []),
- ((Binding.prefix_name "lub_" name, @{thm typedef_lub} OF cpo_thms'), []),
- ((Binding.prefix_name "thelub_" name, @{thm typedef_thelub} OF cpo_thms'), []),
- ((Binding.prefix_name "compact_" name, @{thm typedef_compact} OF cpo_thms'), [])])
- |> snd
- |> Sign.parent_path
- end;
-
- fun make_pcpo UU_mem (type_def, below_def, set_def) theory =
- let
- val UU_mem' = fold_rule (the_list set_def) UU_mem;
- val pcpo_thms = map (Thm.transfer theory) [type_def, below_def, UU_mem'];
- val theory' = theory
- |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
- (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
- val pcpo_thms' = map (Thm.transfer theory') pcpo_thms;
- in
- theory'
- |> Sign.add_path (Binding.name_of name)
- |> PureThy.add_thms
- ([((Binding.suffix_name "_strict" Rep_name, @{thm typedef_Rep_strict} OF pcpo_thms'), []),
- ((Binding.suffix_name "_strict" Abs_name, @{thm typedef_Abs_strict} OF pcpo_thms'), []),
- ((Binding.suffix_name "_strict_iff" Rep_name, @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
- ((Binding.suffix_name "_strict_iff" Abs_name, @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
- ((Binding.suffix_name "_defined" Rep_name, @{thm typedef_Rep_defined} OF pcpo_thms'), []),
- ((Binding.suffix_name "_defined" Abs_name, @{thm typedef_Abs_defined} OF pcpo_thms'), [])])
- |> snd
- |> Sign.parent_path
- end;
-
- fun pcpodef_result UU_mem admissible =
- make_po (Tactic.rtac exI 1 THEN Tactic.rtac UU_mem 1)
- #-> (fn defs => make_cpo admissible defs #> make_pcpo UU_mem defs);
-
- fun cpodef_result nonempty admissible =
- make_po (Tactic.rtac nonempty 1)
- #-> make_cpo admissible;
- in
- if pcpo
- then (goal_UU_mem, goal_admissible, pcpodef_result)
- else (goal_nonempty, goal_admissible, cpodef_result)
- end
- handle ERROR msg =>
- cat_error msg ("The error(s) above occurred in cpodef " ^ quote (Binding.str_of name));
-
-
-(* proof interface *)
-
-local
-
-fun gen_pcpodef_proof prep_term pcpo ((def, name), typ, set, opt_morphs) thy =
- let
- val (goal1, goal2, make_result) =
- prepare_pcpodef prep_term pcpo def name typ set opt_morphs thy;
- fun after_qed [[th1, th2]] = ProofContext.theory (make_result th1 th2);
- in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] (ProofContext.init thy) end;
-
-in
-
-fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term true x;
-fun pcpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term true x;
-
-fun cpodef_proof x = gen_pcpodef_proof Syntax.check_term false x;
-fun cpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term false x;
-
-end;
-
-
-
-(** outer syntax **)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val typedef_proof_decl =
- Scan.optional (P.$$$ "(" |--
- ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
- --| P.$$$ ")") (true, NONE) --
- (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
-
-fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
- (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
- ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs);
-
-val _ =
- OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
- (typedef_proof_decl >>
- (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true)));
-
-val _ =
- OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
- (typedef_proof_decl >>
- (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false)));
-
-end;
-
-end;
--- a/src/Pure/General/position.scala Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Pure/General/position.scala Tue Jun 23 15:48:55 2009 +0100
@@ -11,13 +11,14 @@
object Position {
- private def get_string(name: String, props: Properties) = {
- val value = if (props != null) props.getProperty(name) else null
- if (value != null) Some(value) else None
- }
+ type T = List[(String, String)]
- private def get_int(name: String, props: Properties) = {
- get_string(name, props) match {
+ private def get_string(name: String, pos: T): Option[String] =
+ pos.find(p => p._1 == name).map(_._2)
+
+ private def get_int(name: String, pos: T): Option[Int] =
+ {
+ get_string(name, pos) match {
case None => None
case Some(value) =>
try { Some(Integer.parseInt(value)) }
@@ -25,12 +26,20 @@
}
}
- def line_of(props: Properties) = get_int(Markup.LINE, props)
- def column_of(props: Properties) = get_int(Markup.COLUMN, props)
- def offset_of(props: Properties) = get_int(Markup.OFFSET, props)
- def end_line_of(props: Properties) = get_int(Markup.END_LINE, props)
- def end_column_of(props: Properties) = get_int(Markup.END_COLUMN, props)
- def end_offset_of(props: Properties) = get_int(Markup.END_OFFSET, props)
- def file_of(props: Properties) = get_string(Markup.FILE, props)
- def id_of(props: Properties) = get_string(Markup.ID, props)
+ def line_of(pos: T) = get_int(Markup.LINE, pos)
+ def column_of(pos: T) = get_int(Markup.COLUMN, pos)
+ def offset_of(pos: T) = get_int(Markup.OFFSET, pos)
+ def end_line_of(pos: T) = get_int(Markup.END_LINE, pos)
+ def end_column_of(pos: T) = get_int(Markup.END_COLUMN, pos)
+ def end_offset_of(pos: T) = get_int(Markup.END_OFFSET, pos)
+ def file_of(pos: T) = get_string(Markup.FILE, pos)
+ def id_of(pos: T) = get_string(Markup.ID, pos)
+
+ def offsets_of(pos: T): (Option[Int], Option[Int]) =
+ {
+ val begin = offset_of(pos)
+ val end = end_offset_of(pos)
+ (begin, if (end.isDefined) end else begin.map(_ + 1))
+ }
+
}
--- a/src/Pure/General/scan.scala Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Pure/General/scan.scala Tue Jun 23 15:48:55 2009 +0100
@@ -19,11 +19,8 @@
private case class Tree(val branches: Map[Char, (String, Tree)])
private val empty_tree = Tree(Map())
- private def make(tree: Tree, max: Int): Lexicon =
- new Lexicon {
- override val main_tree = tree
- override val max_entry = max
- }
+ private def make(tree: Tree): Lexicon =
+ new Lexicon { override val main_tree = tree }
val empty: Lexicon = new Lexicon
def apply(strs: String*): Lexicon = (empty /: strs) ((lex, str) => lex + str)
@@ -35,7 +32,6 @@
import Lexicon.Tree
val main_tree: Tree = Lexicon.empty_tree
- val max_entry = -1
/* auxiliary operations */
@@ -74,7 +70,7 @@
override def stringPrefix = "Lexicon"
- override def isEmpty: Boolean = { max_entry < 0 }
+ override def isEmpty: Boolean = { main_tree.branches.isEmpty }
def size: Int = content(main_tree, Nil).length
def elements: Iterator[String] = content(main_tree, Nil).sort(_ <= _).elements
@@ -102,7 +98,7 @@
} else tree
}
if (contains(str)) this
- else Lexicon.make(extend(main_tree, 0), max_entry max str.length)
+ else Lexicon.make(extend(main_tree, 0))
}
def empty[A]: Set[A] = error("Undefined")
--- a/src/Pure/ML-Systems/timing.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Pure/ML-Systems/timing.ML Tue Jun 23 15:48:55 2009 +0100
@@ -19,12 +19,16 @@
val {nongc = {sys = sys2, usr = usr2}, gc = {sys = gc_sys2, usr = gc_usr2}} =
Timer.checkCPUTimes cpu_timer;
+ fun checked x y = Real.checkFloat y handle Overflow => x | Div => x;
+
open Time;
val elapsed = real_time2 - real_time;
val gc = gc_usr2 - gc_usr + gc_sys2 - gc_sys;
val cpu = usr2 - usr + sys2 - sys + gc;
- val gc_percent = Real.fmt (StringCvt.FIX (SOME 1)) (100.0 * toReal gc / toReal cpu);
- val factor = Real.fmt (StringCvt.FIX (SOME 2)) (toReal cpu / toReal elapsed);
+ val gc_percent =
+ Real.fmt (StringCvt.FIX (SOME 1)) (100.0 * checked 0.0 (toReal gc / toReal cpu));
+ val factor =
+ Real.fmt (StringCvt.FIX (SOME 2)) (checked 1.0 (toReal cpu / toReal elapsed));
val message =
(toString elapsed ^ "s elapsed time, " ^ toString cpu ^ "s cpu time, " ^
gc_percent ^ "% GC time, factor " ^ factor) handle Time => "";
--- a/src/Pure/ML/ml_compiler_polyml-5.3.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Pure/ML/ml_compiler_polyml-5.3.ML Tue Jun 23 15:48:55 2009 +0100
@@ -63,6 +63,8 @@
(* eval ML source tokens *)
+val offset_of = the_default 0 o Position.offset_of;
+
fun eval verbose pos toks =
let
val _ = Secure.secure_mltext ();
@@ -80,25 +82,24 @@
val syms = Symbol.explode (ML_Lex.check_content_of tok);
val (ps, _) = fold_map (fn s => fn p => (p, Position.advance s p)) syms
(Position.reset_range (ML_Lex.pos_of tok));
- in ps ~~ map (String.explode o Symbol.esc) syms end);
+ in
+ (ps ~~ syms) |> maps (fn (p, sym) =>
+ map (pair (offset_of p)) (String.explode (Symbol.esc sym)))
+ end);
val end_pos =
if null toks then Position.none
else ML_Lex.end_pos_of (List.last toks);
- val input_buffer = ref (input @ [(end_pos, [#"\n"])]);
+ val input_buffer = ref (input @ [(offset_of end_pos, #"\n")]);
val line = ref (the_default 1 (Position.line_of pos));
- fun get_offset () =
- the_default 0
- (get_first (fn (_, []) => NONE | (p, _) => Position.offset_of p) (! input_buffer));
-
+ fun get_offset () = (case ! input_buffer of [] => 0 | (i, _) :: _ => i);
fun get () =
(case ! input_buffer of
[] => NONE
- | (_, []) :: rest => (input_buffer := rest; get ())
- | (p, c :: cs) :: rest =>
- (input_buffer := (p, cs) :: rest;
+ | (_, c) :: rest =>
+ (input_buffer := rest;
if c = #"\n" then line := ! line + 1 else ();
SOME c));
--- a/src/Pure/System/isabelle_system.scala Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Pure/System/isabelle_system.scala Tue Jun 23 15:48:55 2009 +0100
@@ -98,7 +98,9 @@
if (i <= 0) (entry -> "")
else (entry.substring(0, i) -> entry.substring(i + 1))
}
- Map(entries: _*)
+ Map(entries: _*) +
+ ("HOME" -> java.lang.System.getenv("HOME")) +
+ ("PATH" -> java.lang.System.getenv("PATH"))
}
finally { dump.delete }
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_haskell.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,567 @@
+(* Title: Tools/code/code_haskell.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Serializer for Haskell.
+*)
+
+signature CODE_HASKELL =
+sig
+ val setup: theory -> theory
+end;
+
+structure Code_Haskell : CODE_HASKELL =
+struct
+
+val target = "Haskell";
+
+open Basic_Code_Thingol;
+open Code_Printer;
+
+infixr 5 @@;
+infixr 5 @|;
+
+
+(** Haskell serializer **)
+
+fun pr_haskell_bind pr_term =
+ let
+ fun pr_bind ((NONE, NONE), _) = str "_"
+ | pr_bind ((SOME v, NONE), _) = str v
+ | pr_bind ((NONE, SOME p), _) = p
+ | pr_bind ((SOME v, SOME p), _) = brackets [str v, str "@", p];
+ in gen_pr_bind pr_bind pr_term end;
+
+fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const
+ init_syms deresolve is_cons contr_classparam_typs deriving_show =
+ let
+ val deresolve_base = Long_Name.base_name o deresolve;
+ fun class_name class = case syntax_class class
+ of NONE => deresolve class
+ | SOME class => class;
+ fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
+ of [] => []
+ | classbinds => Pretty.enum "," "(" ")" (
+ map (fn (v, class) =>
+ str (class_name class ^ " " ^ Code_Printer.lookup_var tyvars v)) classbinds)
+ @@ str " => ";
+ fun pr_typforall tyvars vs = case map fst vs
+ of [] => []
+ | vnames => str "forall " :: Pretty.breaks
+ (map (str o Code_Printer.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
+ fun pr_tycoexpr tyvars fxy (tyco, tys) =
+ brackify fxy (str tyco :: map (pr_typ tyvars BR) tys)
+ and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco
+ of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys)
+ | SOME (i, pr) => pr (pr_typ tyvars) fxy tys)
+ | pr_typ tyvars fxy (ITyVar v) = (str o Code_Printer.lookup_var tyvars) v;
+ fun pr_typdecl tyvars (vs, tycoexpr) =
+ Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr);
+ fun pr_typscheme tyvars (vs, ty) =
+ Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty);
+ fun pr_term tyvars thm vars fxy (IConst c) =
+ pr_app tyvars thm vars fxy (c, [])
+ | pr_term tyvars thm vars fxy (t as (t1 `$ t2)) =
+ (case Code_Thingol.unfold_const_app t
+ of SOME app => pr_app tyvars thm vars fxy app
+ | _ =>
+ brackify fxy [
+ pr_term tyvars thm vars NOBR t1,
+ pr_term tyvars thm vars BR t2
+ ])
+ | pr_term tyvars thm vars fxy (IVar v) =
+ (str o Code_Printer.lookup_var vars) v
+ | pr_term tyvars thm vars fxy (t as _ `|=> _) =
+ let
+ val (binds, t') = Code_Thingol.unfold_abs t;
+ fun pr ((v, pat), ty) = pr_bind tyvars thm BR ((SOME v, pat), ty);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end
+ | pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) =
+ (case Code_Thingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+ then pr_case tyvars thm vars fxy cases
+ else pr_app tyvars thm vars fxy c_ts
+ | NONE => pr_case tyvars thm vars fxy cases)
+ and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c
+ of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts
+ | fingerprint => let
+ val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint;
+ val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) =>
+ (not o Code_Thingol.locally_monomorphic) t) ts_fingerprint;
+ fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t
+ | pr_term_anno (t, SOME _) ty =
+ brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty];
+ in
+ if needs_annotation then
+ (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys)
+ else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts
+ end
+ and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const
+ and pr_bind tyvars = pr_haskell_bind (pr_term tyvars)
+ and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, body) = Code_Thingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind tyvars thm BR ((NONE, SOME pat), ty)
+ |>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t])
+ val (ps, vars') = fold_map pr binds vars;
+ in brackify_block fxy (str "let {")
+ ps
+ (concat [str "}", str "in", pr_term tyvars thm vars' NOBR body])
+ end
+ | pr_case tyvars thm vars fxy (((t, ty), clauses as _ :: _), _) =
+ let
+ fun pr (pat, body) =
+ let
+ val (p, vars') = pr_bind tyvars thm NOBR ((NONE, SOME pat), ty) vars;
+ in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end;
+ in brackify_block fxy
+ (concat [str "case", pr_term tyvars thm vars NOBR t, str "of", str "{"])
+ (map pr clauses)
+ (str "}")
+ end
+ | pr_case tyvars thm vars fxy ((_, []), _) =
+ (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""];
+ fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) =
+ let
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ val n = (length o fst o Code_Thingol.unfold_fun) ty;
+ in
+ Pretty.chunks [
+ Pretty.block [
+ (str o suffix " ::" o deresolve_base) name,
+ Pretty.brk 1,
+ pr_typscheme tyvars (vs, ty),
+ str ";"
+ ],
+ concat (
+ (str o deresolve_base) name
+ :: map str (replicate n "_")
+ @ str "="
+ :: str "error"
+ @@ (str o (fn s => s ^ ";") o ML_Syntax.print_string
+ o Long_Name.base_name o Long_Name.qualifier) name
+ )
+ ]
+ end
+ | pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) =
+ let
+ val eqs = filter (snd o snd) raw_eqs;
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ fun pr_eq ((ts, t), (thm, _)) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = init_syms
+ |> Code_Printer.intro_vars consts
+ |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ semicolon (
+ (str o deresolve_base) name
+ :: map (pr_term tyvars thm vars BR) ts
+ @ str "="
+ @@ pr_term tyvars thm vars NOBR t
+ )
+ end;
+ in
+ Pretty.chunks (
+ Pretty.block [
+ (str o suffix " ::" o deresolve_base) name,
+ Pretty.brk 1,
+ pr_typscheme tyvars (vs, ty),
+ str ";"
+ ]
+ :: map pr_eq eqs
+ )
+ end
+ | pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) =
+ let
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ in
+ semicolon [
+ str "data",
+ pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
+ ]
+ end
+ | pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) =
+ let
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ in
+ semicolon (
+ str "newtype"
+ :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
+ :: str "="
+ :: (str o deresolve_base) co
+ :: pr_typ tyvars BR ty
+ :: (if deriving_show name then [str "deriving (Read, Show)"] else [])
+ )
+ end
+ | pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) =
+ let
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ fun pr_co (co, tys) =
+ concat (
+ (str o deresolve_base) co
+ :: map (pr_typ tyvars BR) tys
+ )
+ in
+ semicolon (
+ str "data"
+ :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
+ :: str "="
+ :: pr_co co
+ :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos
+ @ (if deriving_show name then [str "deriving (Read, Show)"] else [])
+ )
+ end
+ | pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) =
+ let
+ val tyvars = Code_Printer.intro_vars [v] init_syms;
+ fun pr_classparam (classparam, ty) =
+ semicolon [
+ (str o deresolve_base) classparam,
+ str "::",
+ pr_typ tyvars NOBR ty
+ ]
+ in
+ Pretty.block_enclose (
+ Pretty.block [
+ str "class ",
+ Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]),
+ str (deresolve_base name ^ " " ^ Code_Printer.lookup_var tyvars v),
+ str " where {"
+ ],
+ str "};"
+ ) (map pr_classparam classparams)
+ end
+ | pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) =
+ let
+ val split_abs_pure = (fn (v, _) `|=> t => SOME (v, t) | _ => NONE);
+ val unfold_abs_pure = Code_Thingol.unfoldr split_abs_pure;
+ val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
+ fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam
+ of NONE => semicolon [
+ (str o deresolve_base) classparam,
+ str "=",
+ pr_app tyvars thm init_syms NOBR (c_inst, [])
+ ]
+ | SOME (k, pr) =>
+ let
+ val (c_inst_name, (_, tys)) = c_inst;
+ val const = if (is_some o syntax_const) c_inst_name
+ then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name;
+ val proto_rhs = Code_Thingol.eta_expand k (c_inst, []);
+ val (vs, rhs) = unfold_abs_pure proto_rhs;
+ val vars = init_syms
+ |> Code_Printer.intro_vars (the_list const)
+ |> Code_Printer.intro_vars vs;
+ val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs;
+ (*dictionaries are not relevant at this late stage*)
+ in
+ semicolon [
+ pr_term tyvars thm vars NOBR lhs,
+ str "=",
+ pr_term tyvars thm vars NOBR rhs
+ ]
+ end;
+ in
+ Pretty.block_enclose (
+ Pretty.block [
+ str "instance ",
+ Pretty.block (pr_typcontext tyvars vs),
+ str (class_name class ^ " "),
+ pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
+ str " where {"
+ ],
+ str "};"
+ ) (map pr_instdef classparam_insts)
+ end;
+ in pr_stmt end;
+
+fun haskell_program_of_program labelled_name module_name module_prefix reserved_names raw_module_alias program =
+ let
+ val module_alias = if is_some module_name then K module_name else raw_module_alias;
+ val reserved_names = Name.make_context reserved_names;
+ val mk_name_module = Code_Printer.mk_name_module reserved_names module_prefix module_alias program;
+ fun add_stmt (name, (stmt, deps)) =
+ let
+ val (module_name, base) = Code_Printer.dest_name name;
+ val module_name' = mk_name_module module_name;
+ val mk_name_stmt = yield_singleton Name.variants;
+ fun add_fun upper (nsp_fun, nsp_typ) =
+ let
+ val (base', nsp_fun') =
+ mk_name_stmt (if upper then Code_Printer.first_upper base else base) nsp_fun
+ in (base', (nsp_fun', nsp_typ)) end;
+ fun add_typ (nsp_fun, nsp_typ) =
+ let
+ val (base', nsp_typ') = mk_name_stmt (Code_Printer.first_upper base) nsp_typ
+ in (base', (nsp_fun, nsp_typ')) end;
+ val add_name = case stmt
+ of Code_Thingol.Fun _ => add_fun false
+ | Code_Thingol.Datatype _ => add_typ
+ | Code_Thingol.Datatypecons _ => add_fun true
+ | Code_Thingol.Class _ => add_typ
+ | Code_Thingol.Classrel _ => pair base
+ | Code_Thingol.Classparam _ => add_fun false
+ | Code_Thingol.Classinst _ => pair base;
+ fun add_stmt' base' = case stmt
+ of Code_Thingol.Datatypecons _ =>
+ cons (name, (Long_Name.append module_name' base', NONE))
+ | Code_Thingol.Classrel _ => I
+ | Code_Thingol.Classparam _ =>
+ cons (name, (Long_Name.append module_name' base', NONE))
+ | _ => cons (name, (Long_Name.append module_name' base', SOME stmt));
+ in
+ Symtab.map_default (module_name', ([], ([], (reserved_names, reserved_names))))
+ (apfst (fold (insert (op = : string * string -> bool)) deps))
+ #> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name'))
+ #-> (fn (base', names) =>
+ (Symtab.map_entry module_name' o apsnd) (fn (stmts, _) =>
+ (add_stmt' base' stmts, names)))
+ end;
+ val hs_program = fold add_stmt (AList.make (fn name =>
+ (Graph.get_node program name, Graph.imm_succs program name))
+ (Graph.strong_conn program |> flat)) Symtab.empty;
+ fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the
+ o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Printer.dest_name) name))) name
+ handle Option => error ("Unknown statement name: " ^ labelled_name name);
+ in (deresolver, hs_program) end;
+
+fun serialize_haskell module_prefix raw_module_name string_classes labelled_name
+ raw_reserved_names includes raw_module_alias
+ syntax_class syntax_tyco syntax_const program cs destination =
+ let
+ val stmt_names = Code_Target.stmt_names_of_destination destination;
+ val module_name = if null stmt_names then raw_module_name else SOME "Code";
+ val reserved_names = fold (insert (op =) o fst) includes raw_reserved_names;
+ val (deresolver, hs_program) = haskell_program_of_program labelled_name
+ module_name module_prefix reserved_names raw_module_alias program;
+ val is_cons = Code_Thingol.is_cons program;
+ val contr_classparam_typs = Code_Thingol.contr_classparam_typs program;
+ fun deriving_show tyco =
+ let
+ fun deriv _ "fun" = false
+ | deriv tycos tyco = member (op =) tycos tyco orelse
+ case try (Graph.get_node program) tyco
+ of SOME (Code_Thingol.Datatype (_, (_, cs))) => forall (deriv' (tyco :: tycos))
+ (maps snd cs)
+ | NONE => true
+ and deriv' tycos (tyco `%% tys) = deriv tycos tyco
+ andalso forall (deriv' tycos) tys
+ | deriv' _ (ITyVar _) = true
+ in deriv [] tyco end;
+ val reserved_names = Code_Printer.make_vars reserved_names;
+ fun pr_stmt qualified = pr_haskell_stmt labelled_name
+ syntax_class syntax_tyco syntax_const reserved_names
+ (if qualified then deresolver else Long_Name.base_name o deresolver)
+ is_cons contr_classparam_typs
+ (if string_classes then deriving_show else K false);
+ fun pr_module name content =
+ (name, Pretty.chunks [
+ str ("module " ^ name ^ " where {"),
+ str "",
+ content,
+ str "",
+ str "}"
+ ]);
+ fun serialize_module1 (module_name', (deps, (stmts, _))) =
+ let
+ val stmt_names = map fst stmts;
+ val deps' = subtract (op =) stmt_names deps
+ |> distinct (op =)
+ |> map_filter (try deresolver);
+ val qualified = is_none module_name andalso
+ map deresolver stmt_names @ deps'
+ |> map Long_Name.base_name
+ |> has_duplicates (op =);
+ val imports = deps'
+ |> map Long_Name.qualifier
+ |> distinct (op =);
+ fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";");
+ val pr_import_module = str o (if qualified
+ then prefix "import qualified "
+ else prefix "import ") o suffix ";";
+ val content = Pretty.chunks (
+ map pr_import_include includes
+ @ map pr_import_module imports
+ @ str ""
+ :: separate (str "") (map_filter
+ (fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt))
+ | (_, (_, NONE)) => NONE) stmts)
+ )
+ in pr_module module_name' content end;
+ fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks (
+ separate (str "") (map_filter
+ (fn (name, (_, SOME stmt)) => if null stmt_names
+ orelse member (op =) stmt_names name
+ then SOME (pr_stmt false (name, stmt))
+ else NONE
+ | (_, (_, NONE)) => NONE) stmts));
+ val serialize_module =
+ if null stmt_names then serialize_module1 else pair "" o serialize_module2;
+ fun check_destination destination =
+ (File.check destination; destination);
+ fun write_module destination (modlname, content) =
+ let
+ val filename = case modlname
+ of "" => Path.explode "Main.hs"
+ | _ => (Path.ext "hs" o Path.explode o implode o separate "/"
+ o Long_Name.explode) modlname;
+ val pathname = Path.append destination filename;
+ val _ = File.mkdir (Path.dir pathname);
+ in File.write pathname
+ ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n"
+ ^ Code_Target.code_of_pretty content)
+ end
+ in
+ Code_Target.mk_serialization target NONE
+ (fn NONE => K () o map (Code_Target.code_writeln o snd) | SOME file => K () o map
+ (write_module (check_destination file)))
+ (rpair [] o cat_lines o map (Code_Target.code_of_pretty o snd))
+ (map (uncurry pr_module) includes
+ @ map serialize_module (Symtab.dest hs_program))
+ destination
+ end;
+
+val literals = let
+ fun char_haskell c =
+ let
+ val s = ML_Syntax.print_char c;
+ in if s = "'" then "\\'" else s end;
+in Literals {
+ literal_char = enclose "'" "'" o char_haskell,
+ literal_string = quote o translate_string char_haskell,
+ literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k
+ else enclose "(" ")" (signed_string_of_int k),
+ literal_list = Pretty.enum "," "[" "]",
+ infix_cons = (5, ":")
+} end;
+
+
+(** optional monad syntax **)
+
+fun pretty_haskell_monad c_bind =
+ let
+ fun dest_bind t1 t2 = case Code_Thingol.split_abs t2
+ of SOME (((v, pat), ty), t') =>
+ SOME ((SOME (((SOME v, pat), ty), true), t1), t')
+ | NONE => NONE;
+ fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) =
+ if c = c_bind_name then dest_bind t1 t2
+ else NONE
+ | dest_monad _ t = case Code_Thingol.split_let t
+ of SOME (((pat, ty), tbind), t') =>
+ SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t')
+ | NONE => NONE;
+ fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name);
+ fun pr_monad pr_bind pr (NONE, t) vars =
+ (semicolon [pr vars NOBR t], vars)
+ | pr_monad pr_bind pr (SOME (bind, true), t) vars = vars
+ |> pr_bind NOBR bind
+ |>> (fn p => semicolon [p, str "<-", pr vars NOBR t])
+ | pr_monad pr_bind pr (SOME (bind, false), t) vars = vars
+ |> pr_bind NOBR bind
+ |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]);
+ fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
+ of SOME (bind, t') => let
+ val (binds, t'') = implode_monad c_bind' t'
+ val (ps, vars') = fold_map (pr_monad (pr_haskell_bind (K pr) thm) pr) (bind :: binds) vars;
+ in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end
+ | NONE => brackify_infix (1, L) fxy
+ [pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2]
+ in (2, ([c_bind], pretty)) end;
+
+fun add_monad target' raw_c_bind thy =
+ let
+ val c_bind = Code.read_const thy raw_c_bind;
+ in if target = target' then
+ thy
+ |> Code_Target.add_syntax_const target c_bind
+ (SOME (pretty_haskell_monad c_bind))
+ else error "Only Haskell target allows for monad syntax" end;
+
+
+(** Isar setup **)
+
+fun isar_seri_haskell module =
+ Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
+ -- Scan.optional (Args.$$$ "string_classes" >> K true) false
+ >> (fn (module_prefix, string_classes) =>
+ serialize_haskell module_prefix module string_classes));
+
+val _ =
+ OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl (
+ OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) =>
+ Toplevel.theory (add_monad target raw_bind))
+ );
+
+val setup =
+ Code_Target.add_target (target, (isar_seri_haskell, literals))
+ #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> fold (Code_Target.add_reserved target) [
+ "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
+ "import", "default", "forall", "let", "in", "class", "qualified", "data",
+ "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
+ ]
+ #> fold (Code_Target.add_reserved target) [
+ "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int",
+ "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded",
+ "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor",
+ "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException",
+ "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException",
+ "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure",
+ "ExitSuccess", "False", "GT", "HeapOverflow",
+ "IOError", "IOException", "IllegalOperation",
+ "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError",
+ "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow",
+ "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError",
+ "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow",
+ "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow",
+ "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all",
+ "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan",
+ "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen",
+ "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break",
+ "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const",
+ "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod",
+ "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem",
+ "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo",
+ "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip",
+ "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational",
+ "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble",
+ "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj",
+ "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head",
+ "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha",
+ "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite",
+ "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'",
+ "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log",
+ "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum",
+ "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem",
+ "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo",
+ "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred",
+ "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn",
+ "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat",
+ "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile",
+ "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn",
+ "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational",
+ "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse",
+ "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence",
+ "sequence_", "show", "showChar", "showException", "showField", "showList",
+ "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand",
+ "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract",
+ "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult",
+ "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry",
+ "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords",
+ "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3"
+ ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*);
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_ml.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,1122 @@
+(* Title: Tools/code/code_ml.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Serializer for SML and OCaml.
+*)
+
+signature CODE_ML =
+sig
+ val eval: string option -> string * (unit -> 'a) option ref
+ -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a
+ val target_Eval: string
+ val setup: theory -> theory
+end;
+
+structure Code_ML : CODE_ML =
+struct
+
+open Basic_Code_Thingol;
+open Code_Printer;
+
+infixr 5 @@;
+infixr 5 @|;
+
+val target_SML = "SML";
+val target_OCaml = "OCaml";
+val target_Eval = "Eval";
+
+datatype ml_stmt =
+ MLExc of string * int
+ | MLVal of string * ((typscheme * iterm) * (thm * bool))
+ | MLFuns of (string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)) list * string list
+ | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list
+ | MLClass of string * (vname * ((class * string) list * (string * itype) list))
+ | MLClassinst of string * ((class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * ((string * const) * (thm * bool)) list));
+
+fun stmt_names_of (MLExc (name, _)) = [name]
+ | stmt_names_of (MLVal (name, _)) = [name]
+ | stmt_names_of (MLFuns (fs, _)) = map fst fs
+ | stmt_names_of (MLDatas ds) = map fst ds
+ | stmt_names_of (MLClass (name, _)) = [name]
+ | stmt_names_of (MLClassinst (name, _)) = [name];
+
+
+(** SML serailizer **)
+
+fun pr_sml_stmt labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons =
+ let
+ fun pr_dicts fxy ds =
+ let
+ fun pr_dictvar (v, (_, 1)) = Code_Printer.first_upper v ^ "_"
+ | pr_dictvar (v, (i, _)) = Code_Printer.first_upper v ^ string_of_int (i+1) ^ "_";
+ fun pr_proj [] p =
+ p
+ | pr_proj [p'] p =
+ brackets [p', p]
+ | pr_proj (ps as _ :: _) p =
+ brackets [Pretty.enum " o" "(" ")" ps, p];
+ fun pr_dict fxy (DictConst (inst, dss)) =
+ brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss)
+ | pr_dict fxy (DictVar (classrels, v)) =
+ pr_proj (map (str o deresolve) classrels) ((str o pr_dictvar) v)
+ in case ds
+ of [] => str "()"
+ | [d] => pr_dict fxy d
+ | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds
+ end;
+ fun pr_tyvar_dicts vs =
+ vs
+ |> map (fn (v, sort) => map_index (fn (i, _) =>
+ DictVar ([], (v, (i, length sort)))) sort)
+ |> map (pr_dicts BR);
+ fun pr_tycoexpr fxy (tyco, tys) =
+ let
+ val tyco' = (str o deresolve) tyco
+ in case map (pr_typ BR) tys
+ of [] => tyco'
+ | [p] => Pretty.block [p, Pretty.brk 1, tyco']
+ | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
+ end
+ and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
+ of NONE => pr_tycoexpr fxy (tyco, tys)
+ | SOME (i, pr) => pr pr_typ fxy tys)
+ | pr_typ fxy (ITyVar v) = str ("'" ^ v);
+ fun pr_term is_closure thm vars fxy (IConst c) =
+ pr_app is_closure thm vars fxy (c, [])
+ | pr_term is_closure thm vars fxy (IVar v) =
+ str (Code_Printer.lookup_var vars v)
+ | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
+ (case Code_Thingol.unfold_const_app t
+ of SOME c_ts => pr_app is_closure thm vars fxy c_ts
+ | NONE => brackify fxy
+ [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
+ | pr_term is_closure thm vars fxy (t as _ `|=> _) =
+ let
+ val (binds, t') = Code_Thingol.unfold_abs t;
+ fun pr ((v, pat), ty) =
+ pr_bind is_closure thm NOBR ((SOME v, pat), ty)
+ #>> (fn p => concat [str "fn", p, str "=>"]);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (ps @ [pr_term is_closure thm vars' NOBR t']) end
+ | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) =
+ (case Code_Thingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+ then pr_case is_closure thm vars fxy cases
+ else pr_app is_closure thm vars fxy c_ts
+ | NONE => pr_case is_closure thm vars fxy cases)
+ and pr_app' is_closure thm vars (app as ((c, ((_, iss), tys)), ts)) =
+ if is_cons c then
+ let
+ val k = length tys
+ in if k < 2 then
+ (str o deresolve) c :: map (pr_term is_closure thm vars BR) ts
+ else if k = length ts then
+ [(str o deresolve) c, Pretty.enum "," "(" ")" (map (pr_term is_closure thm vars NOBR) ts)]
+ else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand k app)] end
+ else if is_closure c
+ then (str o deresolve) c @@ str "()"
+ else
+ (str o deresolve) c
+ :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts
+ and pr_app is_closure thm vars = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
+ syntax_const thm vars
+ and pr_bind' ((NONE, NONE), _) = str "_"
+ | pr_bind' ((SOME v, NONE), _) = str v
+ | pr_bind' ((NONE, SOME p), _) = p
+ | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p]
+ and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
+ and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, body) = Code_Thingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
+ |>> (fn p => semicolon [str "val", p, str "=", pr_term is_closure thm vars NOBR t])
+ val (ps, vars') = fold_map pr binds vars;
+ in
+ Pretty.chunks [
+ [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block,
+ [str ("in"), Pretty.fbrk, pr_term is_closure thm vars' NOBR body] |> Pretty.block,
+ str ("end")
+ ]
+ end
+ | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) =
+ let
+ fun pr delim (pat, body) =
+ let
+ val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
+ in
+ concat [str delim, p, str "=>", pr_term is_closure thm vars' NOBR body]
+ end;
+ in
+ brackets (
+ str "case"
+ :: pr_term is_closure thm vars NOBR t
+ :: pr "of" clause
+ :: map (pr "|") clauses
+ )
+ end
+ | pr_case is_closure thm vars fxy ((_, []), _) =
+ (concat o map str) ["raise", "Fail", "\"empty case\""];
+ fun pr_stmt (MLExc (name, n)) =
+ let
+ val exc_str =
+ (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
+ in
+ (concat o map str) (
+ (if n = 0 then "val" else "fun")
+ :: deresolve name
+ :: replicate n "_"
+ @ "="
+ :: "raise"
+ :: "Fail"
+ @@ exc_str
+ )
+ end
+ | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ (Code_Thingol.fold_constnames (insert (op =)) t []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts;
+ in
+ concat [
+ str "val",
+ (str o deresolve) name,
+ str ":",
+ pr_typ NOBR ty,
+ str "=",
+ pr_term (K false) thm vars NOBR t
+ ]
+ end
+ | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) =
+ let
+ fun pr_funn definer (name, ((vs, ty), eqs as eq :: eqs')) =
+ let
+ val vs_dict = filter_out (null o snd) vs;
+ val shift = if null eqs' then I else
+ map (Pretty.block o single o Pretty.block o single);
+ fun pr_eq definer ((ts, t), (thm, _)) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts
+ |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ concat (
+ str definer
+ :: (str o deresolve) name
+ :: (if member (op =) pseudo_funs name then [str "()"]
+ else pr_tyvar_dicts vs_dict
+ @ map (pr_term (member (op =) pseudo_funs) thm vars BR) ts)
+ @ str "="
+ @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t
+ )
+ end
+ in
+ (Pretty.block o Pretty.fbreaks o shift) (
+ pr_eq definer eq
+ :: map (pr_eq "|") eqs'
+ )
+ end;
+ fun pr_pseudo_fun name = concat [
+ str "val",
+ (str o deresolve) name,
+ str "=",
+ (str o deresolve) name,
+ str "();"
+ ];
+ val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns);
+ val pseudo_ps = map pr_pseudo_fun pseudo_funs;
+ in Pretty.chunks (ps @ Pretty.block ([p, str ";"]) :: pseudo_ps) end
+ | pr_stmt (MLDatas (datas as (data :: datas'))) =
+ let
+ fun pr_co (co, []) =
+ str (deresolve co)
+ | pr_co (co, tys) =
+ concat [
+ str (deresolve co),
+ str "of",
+ Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
+ ];
+ fun pr_data definer (tyco, (vs, [])) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ @@ str "EMPTY__"
+ )
+ | pr_data definer (tyco, (vs, cos)) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ :: separate (str "|") (map pr_co cos)
+ );
+ val (ps, p) = split_last
+ (pr_data "datatype" data :: map (pr_data "and") datas');
+ in Pretty.chunks (ps @| Pretty.block ([p, str ";"])) end
+ | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
+ let
+ val w = Code_Printer.first_upper v ^ "_";
+ fun pr_superclass_field (class, classrel) =
+ (concat o map str) [
+ deresolve classrel, ":", "'" ^ v, deresolve class
+ ];
+ fun pr_classparam_field (classparam, ty) =
+ concat [
+ (str o deresolve) classparam, str ":", pr_typ NOBR ty
+ ];
+ fun pr_classparam_proj (classparam, _) =
+ semicolon [
+ str "fun",
+ (str o deresolve) classparam,
+ Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)],
+ str "=",
+ str ("#" ^ deresolve classparam),
+ str w
+ ];
+ fun pr_superclass_proj (_, classrel) =
+ semicolon [
+ str "fun",
+ (str o deresolve) classrel,
+ Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)],
+ str "=",
+ str ("#" ^ deresolve classrel),
+ str w
+ ];
+ in
+ Pretty.chunks (
+ concat [
+ str ("type '" ^ v),
+ (str o deresolve) class,
+ str "=",
+ Pretty.enum "," "{" "};" (
+ map pr_superclass_field superclasses @ map pr_classparam_field classparams
+ )
+ ]
+ :: map pr_superclass_proj superclasses
+ @ map pr_classparam_proj classparams
+ )
+ end
+ | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) =
+ let
+ fun pr_superclass (_, (classrel, dss)) =
+ concat [
+ (str o Long_Name.base_name o deresolve) classrel,
+ str "=",
+ pr_dicts NOBR [DictConst dss]
+ ];
+ fun pr_classparam ((classparam, c_inst), (thm, _)) =
+ concat [
+ (str o Long_Name.base_name o deresolve) classparam,
+ str "=",
+ pr_app (K false) thm reserved_names NOBR (c_inst, [])
+ ];
+ in
+ semicolon ([
+ str (if null arity then "val" else "fun"),
+ (str o deresolve) inst ] @
+ pr_tyvar_dicts arity @ [
+ str "=",
+ Pretty.enum "," "{" "}"
+ (map pr_superclass superarities @ map pr_classparam classparam_insts),
+ str ":",
+ pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
+ ])
+ end;
+ in pr_stmt end;
+
+fun pr_sml_module name content =
+ Pretty.chunks (
+ str ("structure " ^ name ^ " = ")
+ :: str "struct"
+ :: str ""
+ :: content
+ @ str ""
+ @@ str ("end; (*struct " ^ name ^ "*)")
+ );
+
+val literals_sml = Literals {
+ literal_char = prefix "#" o quote o ML_Syntax.print_char,
+ literal_string = quote o translate_string ML_Syntax.print_char,
+ literal_numeral = fn unbounded => fn k =>
+ if unbounded then "(" ^ string_of_int k ^ " : IntInf.int)"
+ else string_of_int k,
+ literal_list = Pretty.enum "," "[" "]",
+ infix_cons = (7, "::")
+};
+
+
+(** OCaml serializer **)
+
+fun pr_ocaml_stmt labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons =
+ let
+ fun pr_dicts fxy ds =
+ let
+ fun pr_dictvar (v, (_, 1)) = "_" ^ Code_Printer.first_upper v
+ | pr_dictvar (v, (i, _)) = "_" ^ Code_Printer.first_upper v ^ string_of_int (i+1);
+ fun pr_proj ps p =
+ fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p
+ fun pr_dict fxy (DictConst (inst, dss)) =
+ brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss)
+ | pr_dict fxy (DictVar (classrels, v)) =
+ pr_proj (map deresolve classrels) ((str o pr_dictvar) v)
+ in case ds
+ of [] => str "()"
+ | [d] => pr_dict fxy d
+ | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds
+ end;
+ fun pr_tyvar_dicts vs =
+ vs
+ |> map (fn (v, sort) => map_index (fn (i, _) =>
+ DictVar ([], (v, (i, length sort)))) sort)
+ |> map (pr_dicts BR);
+ fun pr_tycoexpr fxy (tyco, tys) =
+ let
+ val tyco' = (str o deresolve) tyco
+ in case map (pr_typ BR) tys
+ of [] => tyco'
+ | [p] => Pretty.block [p, Pretty.brk 1, tyco']
+ | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
+ end
+ and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
+ of NONE => pr_tycoexpr fxy (tyco, tys)
+ | SOME (i, pr) => pr pr_typ fxy tys)
+ | pr_typ fxy (ITyVar v) = str ("'" ^ v);
+ fun pr_term is_closure thm vars fxy (IConst c) =
+ pr_app is_closure thm vars fxy (c, [])
+ | pr_term is_closure thm vars fxy (IVar v) =
+ str (Code_Printer.lookup_var vars v)
+ | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
+ (case Code_Thingol.unfold_const_app t
+ of SOME c_ts => pr_app is_closure thm vars fxy c_ts
+ | NONE =>
+ brackify fxy [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
+ | pr_term is_closure thm vars fxy (t as _ `|=> _) =
+ let
+ val (binds, t') = Code_Thingol.unfold_abs t;
+ fun pr ((v, pat), ty) = pr_bind is_closure thm BR ((SOME v, pat), ty);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (str "fun" :: ps @ str "->" @@ pr_term is_closure thm vars' NOBR t') end
+ | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) = (case Code_Thingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+ then pr_case is_closure thm vars fxy cases
+ else pr_app is_closure thm vars fxy c_ts
+ | NONE => pr_case is_closure thm vars fxy cases)
+ and pr_app' is_closure thm vars (app as ((c, ((_, iss), tys)), ts)) =
+ if is_cons c then
+ if length tys = length ts
+ then case ts
+ of [] => [(str o deresolve) c]
+ | [t] => [(str o deresolve) c, pr_term is_closure thm vars BR t]
+ | _ => [(str o deresolve) c, Pretty.enum "," "(" ")"
+ (map (pr_term is_closure thm vars NOBR) ts)]
+ else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand (length tys) app)]
+ else if is_closure c
+ then (str o deresolve) c @@ str "()"
+ else (str o deresolve) c
+ :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts)
+ and pr_app is_closure = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
+ syntax_const
+ and pr_bind' ((NONE, NONE), _) = str "_"
+ | pr_bind' ((SOME v, NONE), _) = str v
+ | pr_bind' ((NONE, SOME p), _) = p
+ | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v]
+ and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
+ and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, body) = Code_Thingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
+ |>> (fn p => concat
+ [str "let", p, str "=", pr_term is_closure thm vars NOBR t, str "in"])
+ val (ps, vars') = fold_map pr binds vars;
+ in
+ brackify_block fxy (Pretty.chunks ps) []
+ (pr_term is_closure thm vars' NOBR body)
+ end
+ | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) =
+ let
+ fun pr delim (pat, body) =
+ let
+ val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
+ in concat [str delim, p, str "->", pr_term is_closure thm vars' NOBR body] end;
+ in
+ brackets (
+ str "match"
+ :: pr_term is_closure thm vars NOBR t
+ :: pr "with" clause
+ :: map (pr "|") clauses
+ )
+ end
+ | pr_case is_closure thm vars fxy ((_, []), _) =
+ (concat o map str) ["failwith", "\"empty case\""];
+ fun fish_params vars eqs =
+ let
+ fun fish_param _ (w as SOME _) = w
+ | fish_param (IVar v) NONE = SOME v
+ | fish_param _ NONE = NONE;
+ fun fillup_param _ (_, SOME v) = v
+ | fillup_param x (i, NONE) = x ^ string_of_int i;
+ val fished1 = fold (map2 fish_param) eqs (replicate (length (hd eqs)) NONE);
+ val x = Name.variant (map_filter I fished1) "x";
+ val fished2 = map_index (fillup_param x) fished1;
+ val (fished3, _) = Name.variants fished2 Name.context;
+ val vars' = Code_Printer.intro_vars fished3 vars;
+ in map (Code_Printer.lookup_var vars') fished3 end;
+ fun pr_stmt (MLExc (name, n)) =
+ let
+ val exc_str =
+ (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
+ in
+ (concat o map str) (
+ "let"
+ :: deresolve name
+ :: replicate n "_"
+ @ "="
+ :: "failwith"
+ @@ exc_str
+ )
+ end
+ | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ (Code_Thingol.fold_constnames (insert (op =)) t []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts;
+ in
+ concat [
+ str "let",
+ (str o deresolve) name,
+ str ":",
+ pr_typ NOBR ty,
+ str "=",
+ pr_term (K false) thm vars NOBR t
+ ]
+ end
+ | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) =
+ let
+ fun pr_eq ((ts, t), (thm, _)) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts
+ |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in concat [
+ (Pretty.block o Pretty.commas)
+ (map (pr_term (member (op =) pseudo_funs) thm vars NOBR) ts),
+ str "->",
+ pr_term (member (op =) pseudo_funs) thm vars NOBR t
+ ] end;
+ fun pr_eqs is_pseudo [((ts, t), (thm, _))] =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts
+ |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ concat (
+ (if is_pseudo then [str "()"]
+ else map (pr_term (member (op =) pseudo_funs) thm vars BR) ts)
+ @ str "="
+ @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t
+ )
+ end
+ | pr_eqs _ (eqs as (eq as (([_], _), _)) :: eqs') =
+ Pretty.block (
+ str "="
+ :: Pretty.brk 1
+ :: str "function"
+ :: Pretty.brk 1
+ :: pr_eq eq
+ :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1]
+ o single o pr_eq) eqs'
+ )
+ | pr_eqs _ (eqs as eq :: eqs') =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c)
+ ((fold o Code_Thingol.fold_constnames)
+ (insert (op =)) (map (snd o fst) eqs) []);
+ val vars = reserved_names
+ |> Code_Printer.intro_vars consts;
+ val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs;
+ in
+ Pretty.block (
+ Pretty.breaks dummy_parms
+ @ Pretty.brk 1
+ :: str "="
+ :: Pretty.brk 1
+ :: str "match"
+ :: Pretty.brk 1
+ :: (Pretty.block o Pretty.commas) dummy_parms
+ :: Pretty.brk 1
+ :: str "with"
+ :: Pretty.brk 1
+ :: pr_eq eq
+ :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1]
+ o single o pr_eq) eqs'
+ )
+ end;
+ fun pr_funn definer (name, ((vs, ty), eqs)) =
+ concat (
+ str definer
+ :: (str o deresolve) name
+ :: pr_tyvar_dicts (filter_out (null o snd) vs)
+ @| pr_eqs (member (op =) pseudo_funs name) eqs
+ );
+ fun pr_pseudo_fun name = concat [
+ str "let",
+ (str o deresolve) name,
+ str "=",
+ (str o deresolve) name,
+ str "();;"
+ ];
+ val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns);
+ val (ps, p) = split_last
+ (pr_funn "let rec" funn :: map (pr_funn "and") funns);
+ val pseudo_ps = map pr_pseudo_fun pseudo_funs;
+ in Pretty.chunks (ps @ Pretty.block ([p, str ";;"]) :: pseudo_ps) end
+ | pr_stmt (MLDatas (datas as (data :: datas'))) =
+ let
+ fun pr_co (co, []) =
+ str (deresolve co)
+ | pr_co (co, tys) =
+ concat [
+ str (deresolve co),
+ str "of",
+ Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
+ ];
+ fun pr_data definer (tyco, (vs, [])) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ @@ str "EMPTY_"
+ )
+ | pr_data definer (tyco, (vs, cos)) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ :: separate (str "|") (map pr_co cos)
+ );
+ val (ps, p) = split_last
+ (pr_data "type" data :: map (pr_data "and") datas');
+ in Pretty.chunks (ps @| Pretty.block ([p, str ";;"])) end
+ | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
+ let
+ val w = "_" ^ Code_Printer.first_upper v;
+ fun pr_superclass_field (class, classrel) =
+ (concat o map str) [
+ deresolve classrel, ":", "'" ^ v, deresolve class
+ ];
+ fun pr_classparam_field (classparam, ty) =
+ concat [
+ (str o deresolve) classparam, str ":", pr_typ NOBR ty
+ ];
+ fun pr_classparam_proj (classparam, _) =
+ concat [
+ str "let",
+ (str o deresolve) classparam,
+ str w,
+ str "=",
+ str (w ^ "." ^ deresolve classparam ^ ";;")
+ ];
+ in Pretty.chunks (
+ concat [
+ str ("type '" ^ v),
+ (str o deresolve) class,
+ str "=",
+ enum_default "unit;;" ";" "{" "};;" (
+ map pr_superclass_field superclasses
+ @ map pr_classparam_field classparams
+ )
+ ]
+ :: map pr_classparam_proj classparams
+ ) end
+ | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) =
+ let
+ fun pr_superclass (_, (classrel, dss)) =
+ concat [
+ (str o deresolve) classrel,
+ str "=",
+ pr_dicts NOBR [DictConst dss]
+ ];
+ fun pr_classparam_inst ((classparam, c_inst), (thm, _)) =
+ concat [
+ (str o deresolve) classparam,
+ str "=",
+ pr_app (K false) thm reserved_names NOBR (c_inst, [])
+ ];
+ in
+ concat (
+ str "let"
+ :: (str o deresolve) inst
+ :: pr_tyvar_dicts arity
+ @ str "="
+ @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [
+ enum_default "()" ";" "{" "}" (map pr_superclass superarities
+ @ map pr_classparam_inst classparam_insts),
+ str ":",
+ pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
+ ]
+ )
+ end;
+ in pr_stmt end;
+
+fun pr_ocaml_module name content =
+ Pretty.chunks (
+ str ("module " ^ name ^ " = ")
+ :: str "struct"
+ :: str ""
+ :: content
+ @ str ""
+ @@ str ("end;; (*struct " ^ name ^ "*)")
+ );
+
+val literals_ocaml = let
+ fun chr i =
+ let
+ val xs = string_of_int i;
+ val ys = replicate_string (3 - length (explode xs)) "0";
+ in "\\" ^ ys ^ xs end;
+ fun char_ocaml c =
+ let
+ val i = ord c;
+ val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126
+ then chr i else c
+ in s end;
+ fun bignum_ocaml k = if k <= 1073741823
+ then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")"
+ else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")"
+in Literals {
+ literal_char = enclose "'" "'" o char_ocaml,
+ literal_string = quote o translate_string char_ocaml,
+ literal_numeral = fn unbounded => fn k => if k >= 0 then
+ if unbounded then bignum_ocaml k
+ else string_of_int k
+ else
+ if unbounded then "(Big_int.minus_big_int " ^ bignum_ocaml (~ k) ^ ")"
+ else (enclose "(" ")" o prefix "-" o string_of_int o op ~) k,
+ literal_list = Pretty.enum ";" "[" "]",
+ infix_cons = (6, "::")
+} end;
+
+
+
+(** SML/OCaml generic part **)
+
+local
+
+datatype ml_node =
+ Dummy of string
+ | Stmt of string * ml_stmt
+ | Module of string * ((Name.context * Name.context) * ml_node Graph.T);
+
+in
+
+fun ml_node_of_program labelled_name module_name reserved_names raw_module_alias program =
+ let
+ val module_alias = if is_some module_name then K module_name else raw_module_alias;
+ val reserved_names = Name.make_context reserved_names;
+ val empty_module = ((reserved_names, reserved_names), Graph.empty);
+ fun map_node [] f = f
+ | map_node (m::ms) f =
+ Graph.default_node (m, Module (m, empty_module))
+ #> Graph.map_node m (fn (Module (module_name, (nsp, nodes))) =>
+ Module (module_name, (nsp, map_node ms f nodes)));
+ fun map_nsp_yield [] f (nsp, nodes) =
+ let
+ val (x, nsp') = f nsp
+ in (x, (nsp', nodes)) end
+ | map_nsp_yield (m::ms) f (nsp, nodes) =
+ let
+ val (x, nodes') =
+ nodes
+ |> Graph.default_node (m, Module (m, empty_module))
+ |> Graph.map_node_yield m (fn Module (d_module_name, nsp_nodes) =>
+ let
+ val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes
+ in (x, Module (d_module_name, nsp_nodes')) end)
+ in (x, (nsp, nodes')) end;
+ fun map_nsp_fun_yield f (nsp_fun, nsp_typ) =
+ let
+ val (x, nsp_fun') = f nsp_fun
+ in (x, (nsp_fun', nsp_typ)) end;
+ fun map_nsp_typ_yield f (nsp_fun, nsp_typ) =
+ let
+ val (x, nsp_typ') = f nsp_typ
+ in (x, (nsp_fun, nsp_typ')) end;
+ val mk_name_module = Code_Printer.mk_name_module reserved_names NONE module_alias program;
+ fun mk_name_stmt upper name nsp =
+ let
+ val (_, base) = Code_Printer.dest_name name;
+ val base' = if upper then Code_Printer.first_upper base else base;
+ val ([base''], nsp') = Name.variants [base'] nsp;
+ in (base'', nsp') end;
+ fun rearrange_fun name (tysm as (vs, ty), raw_eqs) =
+ let
+ val eqs = filter (snd o snd) raw_eqs;
+ val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs
+ of [(([], t), thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty
+ then ([(([IVar "x"], t `$ IVar "x"), thm)], false)
+ else (eqs, not (Code_Thingol.fold_constnames
+ (fn name' => fn b => b orelse name = name') t false))
+ | _ => (eqs, false)
+ else (eqs, false)
+ in ((name, (tysm, eqs')), is_value) end;
+ fun check_kind [((name, (tysm, [(([], t), thm)])), true)] = MLVal (name, ((tysm, t), thm))
+ | check_kind [((name, ((vs, ty), [])), _)] =
+ MLExc (name, (length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty)
+ | check_kind funns =
+ MLFuns (map fst funns, map_filter
+ (fn ((name, ((vs, _), [(([], _), _)])), _) =>
+ if null (filter_out (null o snd) vs) then SOME name else NONE
+ | _ => NONE) funns);
+ fun add_funs stmts = fold_map
+ (fn (name, Code_Thingol.Fun (_, stmt)) =>
+ map_nsp_fun_yield (mk_name_stmt false name)
+ #>> rpair (rearrange_fun name stmt)
+ | (name, _) =>
+ error ("Function block containing illegal statement: " ^ labelled_name name)
+ ) stmts
+ #>> (split_list #> apsnd check_kind);
+ fun add_datatypes stmts =
+ fold_map
+ (fn (name, Code_Thingol.Datatype (_, stmt)) =>
+ map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
+ | (name, Code_Thingol.Datatypecons _) =>
+ map_nsp_fun_yield (mk_name_stmt true name) #>> rpair NONE
+ | (name, _) =>
+ error ("Datatype block containing illegal statement: " ^ labelled_name name)
+ ) stmts
+ #>> (split_list #> apsnd (map_filter I
+ #> (fn [] => error ("Datatype block without data statement: "
+ ^ (commas o map (labelled_name o fst)) stmts)
+ | stmts => MLDatas stmts)));
+ fun add_class stmts =
+ fold_map
+ (fn (name, Code_Thingol.Class (_, stmt)) =>
+ map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
+ | (name, Code_Thingol.Classrel _) =>
+ map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
+ | (name, Code_Thingol.Classparam _) =>
+ map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
+ | (name, _) =>
+ error ("Class block containing illegal statement: " ^ labelled_name name)
+ ) stmts
+ #>> (split_list #> apsnd (map_filter I
+ #> (fn [] => error ("Class block without class statement: "
+ ^ (commas o map (labelled_name o fst)) stmts)
+ | [stmt] => MLClass stmt)));
+ fun add_inst [(name, Code_Thingol.Classinst stmt)] =
+ map_nsp_fun_yield (mk_name_stmt false name)
+ #>> (fn base => ([base], MLClassinst (name, stmt)));
+ fun add_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) =
+ add_funs stmts
+ | add_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) =
+ add_datatypes stmts
+ | add_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) =
+ add_datatypes stmts
+ | add_stmts ((stmts as (_, Code_Thingol.Class _)::_)) =
+ add_class stmts
+ | add_stmts ((stmts as (_, Code_Thingol.Classrel _)::_)) =
+ add_class stmts
+ | add_stmts ((stmts as (_, Code_Thingol.Classparam _)::_)) =
+ add_class stmts
+ | add_stmts ((stmts as [(_, Code_Thingol.Classinst _)])) =
+ add_inst stmts
+ | add_stmts stmts = error ("Illegal mutual dependencies: " ^
+ (commas o map (labelled_name o fst)) stmts);
+ fun add_stmts' stmts nsp_nodes =
+ let
+ val names as (name :: names') = map fst stmts;
+ val deps =
+ []
+ |> fold (fold (insert (op =)) o Graph.imm_succs program) names
+ |> subtract (op =) names;
+ val (module_names, _) = (split_list o map Code_Printer.dest_name) names;
+ val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
+ handle Empty =>
+ error ("Different namespace prefixes for mutual dependencies:\n"
+ ^ commas (map labelled_name names)
+ ^ "\n"
+ ^ commas module_names);
+ val module_name_path = Long_Name.explode module_name;
+ fun add_dep name name' =
+ let
+ val module_name' = (mk_name_module o fst o Code_Printer.dest_name) name';
+ in if module_name = module_name' then
+ map_node module_name_path (Graph.add_edge (name, name'))
+ else let
+ val (common, (diff1 :: _, diff2 :: _)) = chop_prefix (op =)
+ (module_name_path, Long_Name.explode module_name');
+ in
+ map_node common
+ (fn node => Graph.add_edge_acyclic (diff1, diff2) node
+ handle Graph.CYCLES _ => error ("Dependency "
+ ^ quote name ^ " -> " ^ quote name'
+ ^ " would result in module dependency cycle"))
+ end end;
+ in
+ nsp_nodes
+ |> map_nsp_yield module_name_path (add_stmts stmts)
+ |-> (fn (base' :: bases', stmt') =>
+ apsnd (map_node module_name_path (Graph.new_node (name, (Stmt (base', stmt')))
+ #> fold2 (fn name' => fn base' =>
+ Graph.new_node (name', (Dummy base'))) names' bases')))
+ |> apsnd (fold (fn name => fold (add_dep name) deps) names)
+ |> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names)
+ end;
+ val (_, nodes) = empty_module
+ |> fold add_stmts' (map (AList.make (Graph.get_node program))
+ (rev (Graph.strong_conn program)));
+ fun deresolver prefix name =
+ let
+ val module_name = (fst o Code_Printer.dest_name) name;
+ val module_name' = (Long_Name.explode o mk_name_module) module_name;
+ val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name');
+ val stmt_name =
+ nodes
+ |> fold (fn name => fn node => case Graph.get_node node name
+ of Module (_, (_, node)) => node) module_name'
+ |> (fn node => case Graph.get_node node name of Stmt (stmt_name, _) => stmt_name
+ | Dummy stmt_name => stmt_name);
+ in
+ Long_Name.implode (remainder @ [stmt_name])
+ end handle Graph.UNDEF _ =>
+ error ("Unknown statement name: " ^ labelled_name name);
+ in (deresolver, nodes) end;
+
+fun serialize_ml target compile pr_module pr_stmt raw_module_name labelled_name reserved_names includes raw_module_alias
+ _ syntax_tyco syntax_const program stmt_names destination =
+ let
+ val is_cons = Code_Thingol.is_cons program;
+ val present_stmt_names = Code_Target.stmt_names_of_destination destination;
+ val is_present = not (null present_stmt_names);
+ val module_name = if is_present then SOME "Code" else raw_module_name;
+ val (deresolver, nodes) = ml_node_of_program labelled_name module_name
+ reserved_names raw_module_alias program;
+ val reserved_names = Code_Printer.make_vars reserved_names;
+ fun pr_node prefix (Dummy _) =
+ NONE
+ | pr_node prefix (Stmt (_, stmt)) = if is_present andalso
+ (null o filter (member (op =) present_stmt_names) o stmt_names_of) stmt
+ then NONE
+ else SOME
+ (pr_stmt labelled_name syntax_tyco syntax_const reserved_names
+ (deresolver prefix) is_cons stmt)
+ | pr_node prefix (Module (module_name, (_, nodes))) =
+ separate (str "")
+ ((map_filter (pr_node (prefix @ [module_name]) o Graph.get_node nodes)
+ o rev o flat o Graph.strong_conn) nodes)
+ |> (if is_present then Pretty.chunks else pr_module module_name)
+ |> SOME;
+ val stmt_names' = (map o try)
+ (deresolver (if is_some module_name then the_list module_name else [])) stmt_names;
+ val p = Pretty.chunks (separate (str "") (map snd includes @ (map_filter
+ (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes));
+ in
+ Code_Target.mk_serialization target
+ (case compile of SOME compile => SOME (compile o Code_Target.code_of_pretty) | NONE => NONE)
+ (fn NONE => Code_Target.code_writeln | SOME file => File.write file o Code_Target.code_of_pretty)
+ (rpair stmt_names' o Code_Target.code_of_pretty) p destination
+ end;
+
+end; (*local*)
+
+
+(** ML (system language) code for evaluation and instrumentalization **)
+
+fun eval_code_of some_target thy = Code_Target.serialize_custom thy (the_default target_Eval some_target,
+ (fn _ => fn [] => serialize_ml target_SML (SOME (K ())) (K Pretty.chunks) pr_sml_stmt (SOME ""),
+ literals_sml));
+
+
+(* evaluation *)
+
+fun eval some_target reff postproc thy t args =
+ let
+ val ctxt = ProofContext.init thy;
+ fun evaluator naming program ((_, (_, ty)), t) deps =
+ let
+ val _ = if Code_Thingol.contains_dictvar t then
+ error "Term to be evaluated contains free dictionaries" else ();
+ val value_name = "Value.VALUE.value"
+ val program' = program
+ |> Graph.new_node (value_name,
+ Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))])))
+ |> fold (curry Graph.add_edge value_name) deps;
+ val (value_code, [SOME value_name']) = eval_code_of some_target thy naming program' [value_name];
+ val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name'
+ ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
+ in ML_Context.evaluate ctxt false reff sml_code end;
+ in Code_Thingol.eval thy I postproc evaluator t end;
+
+
+(* instrumentalization by antiquotation *)
+
+local
+
+structure CodeAntiqData = ProofDataFun
+(
+ type T = (string list * string list) * (bool * (string
+ * (string * ((string * string) list * (string * string) list)) lazy));
+ fun init _ = (([], []), (true, ("", Lazy.value ("", ([], [])))));
+);
+
+val is_first_occ = fst o snd o CodeAntiqData.get;
+
+fun delayed_code thy tycos consts () =
+ let
+ val (consts', (naming, program)) = Code_Thingol.consts_program thy consts;
+ val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos;
+ val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos');
+ val (consts'', tycos'') = chop (length consts') target_names;
+ val consts_map = map2 (fn const => fn NONE =>
+ error ("Constant " ^ (quote o Code.string_of_const thy) const
+ ^ "\nhas a user-defined serialization")
+ | SOME const'' => (const, const'')) consts consts''
+ val tycos_map = map2 (fn tyco => fn NONE =>
+ error ("Type " ^ (quote o Sign.extern_type thy) tyco
+ ^ "\nhas a user-defined serialization")
+ | SOME tyco'' => (tyco, tyco'')) tycos tycos'';
+ in (ml_code, (tycos_map, consts_map)) end;
+
+fun register_code new_tycos new_consts ctxt =
+ let
+ val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt;
+ val tycos' = fold (insert (op =)) new_tycos tycos;
+ val consts' = fold (insert (op =)) new_consts consts;
+ val (struct_name', ctxt') = if struct_name = ""
+ then ML_Antiquote.variant "Code" ctxt
+ else (struct_name, ctxt);
+ val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) tycos' consts');
+ in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end;
+
+fun register_const const = register_code [] [const];
+
+fun register_datatype tyco constrs = register_code [tyco] constrs;
+
+fun print_const const all_struct_name tycos_map consts_map =
+ (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const;
+
+fun print_datatype tyco constrs all_struct_name tycos_map consts_map =
+ let
+ val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode;
+ fun check_base name name'' =
+ if upperize (Long_Name.base_name name) = upperize name''
+ then () else error ("Name as printed " ^ quote name''
+ ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry.");
+ val tyco'' = (the o AList.lookup (op =) tycos_map) tyco;
+ val constrs'' = map (the o AList.lookup (op =) consts_map) constrs;
+ val _ = check_base tyco tyco'';
+ val _ = map2 check_base constrs constrs'';
+ in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end;
+
+fun print_code struct_name is_first print_it ctxt =
+ let
+ val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt;
+ val (raw_ml_code, (tycos_map, consts_map)) = Lazy.force acc_code;
+ val ml_code = if is_first then "\nstructure " ^ struct_code_name
+ ^ " =\nstruct\n\n" ^ raw_ml_code ^ "\nend;\n\n"
+ else "";
+ val all_struct_name = Long_Name.append struct_name struct_code_name;
+ in (ml_code, print_it all_struct_name tycos_map consts_map) end;
+
+in
+
+fun ml_code_antiq raw_const {struct_name, background} =
+ let
+ val const = Code.check_const (ProofContext.theory_of background) raw_const;
+ val is_first = is_first_occ background;
+ val background' = register_const const background;
+ in (print_code struct_name is_first (print_const const), background') end;
+
+fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} =
+ let
+ val thy = ProofContext.theory_of background;
+ val tyco = Sign.intern_type thy raw_tyco;
+ val constrs = map (Code.check_const thy) raw_constrs;
+ val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
+ val _ = if gen_eq_set (op =) (constrs, constrs') then ()
+ else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
+ val is_first = is_first_occ background;
+ val background' = register_datatype tyco constrs background;
+ in (print_code struct_name is_first (print_datatype tyco constrs), background') end;
+
+end; (*local*)
+
+
+(** Isar setup **)
+
+val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq);
+val _ = ML_Context.add_antiq "code_datatype" (fn _ =>
+ (Args.tyname --| Scan.lift (Args.$$$ "=")
+ -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term)))
+ >> ml_code_datatype_antiq);
+
+fun isar_seri_sml module_name =
+ Code_Target.parse_args (Scan.succeed ())
+ #> (fn () => serialize_ml target_SML
+ (SOME (use_text ML_Env.local_context (1, "generated code") false))
+ pr_sml_module pr_sml_stmt module_name);
+
+fun isar_seri_ocaml module_name =
+ Code_Target.parse_args (Scan.succeed ())
+ #> (fn () => serialize_ml target_OCaml NONE
+ pr_ocaml_module pr_ocaml_stmt module_name);
+
+val setup =
+ Code_Target.add_target (target_SML, (isar_seri_sml, literals_sml))
+ #> Code_Target.add_target (target_OCaml, (isar_seri_ocaml, literals_ocaml))
+ #> Code_Target.extend_target (target_Eval, (target_SML, K I))
+ #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> Code_Target.add_syntax_tyco target_OCaml "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names
+ #> fold (Code_Target.add_reserved target_SML)
+ ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)]
+ #> fold (Code_Target.add_reserved target_OCaml) [
+ "and", "as", "assert", "begin", "class",
+ "constraint", "do", "done", "downto", "else", "end", "exception",
+ "external", "false", "for", "fun", "function", "functor", "if",
+ "in", "include", "inherit", "initializer", "lazy", "let", "match", "method",
+ "module", "mutable", "new", "object", "of", "open", "or", "private", "rec",
+ "sig", "struct", "then", "to", "true", "try", "type", "val",
+ "virtual", "when", "while", "with"
+ ]
+ #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod"];
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_preproc.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,515 @@
+(* Title: Tools/code/code_preproc.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Preprocessing code equations into a well-sorted system
+in a graph with explicit dependencies.
+*)
+
+signature CODE_PREPROC =
+sig
+ val map_pre: (simpset -> simpset) -> theory -> theory
+ val map_post: (simpset -> simpset) -> theory -> theory
+ val add_inline: thm -> theory -> theory
+ val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
+ val del_functrans: string -> theory -> theory
+ val simple_functrans: (theory -> thm list -> thm list option)
+ -> theory -> (thm * bool) list -> (thm * bool) list option
+ val print_codeproc: theory -> unit
+
+ type code_algebra
+ type code_graph
+ val eqns: code_graph -> string -> (thm * bool) list
+ val typ: code_graph -> string -> (string * sort) list * typ
+ val all: code_graph -> string list
+ val pretty: theory -> code_graph -> Pretty.T
+ val obtain: theory -> string list -> term list -> code_algebra * code_graph
+ val eval_conv: theory -> (sort -> sort)
+ -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
+ val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
+ -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
+
+ val setup: theory -> theory
+end
+
+structure Code_Preproc : CODE_PREPROC =
+struct
+
+(** preprocessor administration **)
+
+(* theory data *)
+
+datatype thmproc = Thmproc of {
+ pre: simpset,
+ post: simpset,
+ functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
+};
+
+fun make_thmproc ((pre, post), functrans) =
+ Thmproc { pre = pre, post = post, functrans = functrans };
+fun map_thmproc f (Thmproc { pre, post, functrans }) =
+ make_thmproc (f ((pre, post), functrans));
+fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
+ Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
+ let
+ val pre = Simplifier.merge_ss (pre1, pre2);
+ val post = Simplifier.merge_ss (post1, post2);
+ val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
+ in make_thmproc ((pre, post), functrans) end;
+
+structure Code_Preproc_Data = TheoryDataFun
+(
+ type T = thmproc;
+ val empty = make_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
+ fun copy spec = spec;
+ val extend = copy;
+ fun merge pp = merge_thmproc;
+);
+
+fun the_thmproc thy = case Code_Preproc_Data.get thy
+ of Thmproc x => x;
+
+fun delete_force msg key xs =
+ if AList.defined (op =) xs key then AList.delete (op =) key xs
+ else error ("No such " ^ msg ^ ": " ^ quote key);
+
+fun map_data f thy =
+ thy
+ |> Code.purge_data
+ |> (Code_Preproc_Data.map o map_thmproc) f;
+
+val map_pre = map_data o apfst o apfst;
+val map_post = map_data o apfst o apsnd;
+
+val add_inline = map_pre o MetaSimplifier.add_simp;
+val del_inline = map_pre o MetaSimplifier.del_simp;
+val add_post = map_post o MetaSimplifier.add_simp;
+val del_post = map_post o MetaSimplifier.del_simp;
+
+fun add_functrans (name, f) = (map_data o apsnd)
+ (AList.update (op =) (name, (serial (), f)));
+
+fun del_functrans name = (map_data o apsnd)
+ (delete_force "function transformer" name);
+
+
+(* post- and preprocessing *)
+
+fun apply_functrans thy c _ [] = []
+ | apply_functrans thy c [] eqns = eqns
+ | apply_functrans thy c functrans eqns = eqns
+ |> perhaps (perhaps_loop (perhaps_apply functrans))
+ |> Code.assert_eqns_const thy c;
+
+fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
+
+fun term_of_conv thy f =
+ Thm.cterm_of thy
+ #> f
+ #> Thm.prop_of
+ #> Logic.dest_equals
+ #> snd;
+
+fun preprocess thy c eqns =
+ let
+ val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
+ val functrans = (map (fn (_, (_, f)) => f thy) o #functrans
+ o the_thmproc) thy;
+ in
+ eqns
+ |> apply_functrans thy c functrans
+ |> (map o apfst) (Code.rewrite_eqn pre)
+ |> (map o apfst) (AxClass.unoverload thy)
+ |> map (Code.assert_eqn thy)
+ |> burrow_fst (Code.norm_args thy)
+ |> burrow_fst (Code.norm_varnames thy)
+ end;
+
+fun preprocess_conv thy ct =
+ let
+ val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
+ in
+ ct
+ |> Simplifier.rewrite pre
+ |> rhs_conv (AxClass.unoverload_conv thy)
+ end;
+
+fun postprocess_conv thy ct =
+ let
+ val post = (Simplifier.theory_context thy o #post o the_thmproc) thy;
+ in
+ ct
+ |> AxClass.overload_conv thy
+ |> rhs_conv (Simplifier.rewrite post)
+ end;
+
+fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
+
+fun print_codeproc thy =
+ let
+ val ctxt = ProofContext.init thy;
+ val pre = (#pre o the_thmproc) thy;
+ val post = (#post o the_thmproc) thy;
+ val functrans = (map fst o #functrans o the_thmproc) thy;
+ in
+ (Pretty.writeln o Pretty.chunks) [
+ Pretty.block [
+ Pretty.str "preprocessing simpset:",
+ Pretty.fbrk,
+ Simplifier.pretty_ss ctxt pre
+ ],
+ Pretty.block [
+ Pretty.str "postprocessing simpset:",
+ Pretty.fbrk,
+ Simplifier.pretty_ss ctxt post
+ ],
+ Pretty.block (
+ Pretty.str "function transformers:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map Pretty.str) functrans
+ )
+ ]
+ end;
+
+fun simple_functrans f thy eqns = case f thy (map fst eqns)
+ of SOME thms' => SOME (map (rpair (forall snd eqns)) thms')
+ | NONE => NONE;
+
+
+(** sort algebra and code equation graph types **)
+
+type code_algebra = (sort -> sort) * Sorts.algebra;
+type code_graph = (((string * sort) list * typ) * (thm * bool) list) Graph.T;
+
+fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr);
+fun typ eqngr = fst o Graph.get_node eqngr;
+fun all eqngr = Graph.keys eqngr;
+
+fun pretty thy eqngr =
+ AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
+ |> (map o apfst) (Code.string_of_const thy)
+ |> sort (string_ord o pairself fst)
+ |> map (fn (s, thms) =>
+ (Pretty.block o Pretty.fbreaks) (
+ Pretty.str s
+ :: map (Display.pretty_thm o fst) thms
+ ))
+ |> Pretty.chunks;
+
+
+(** the Waisenhaus algorithm **)
+
+(* auxiliary *)
+
+fun is_proper_class thy = can (AxClass.get_info thy);
+
+fun complete_proper_sort thy =
+ Sign.complete_sort thy #> filter (is_proper_class thy);
+
+fun inst_params thy tyco =
+ map (fn (c, _) => AxClass.param_of_inst thy (c, tyco))
+ o maps (#params o AxClass.get_info thy);
+
+fun consts_of thy eqns = [] |> (fold o fold o fold_aterms)
+ (fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I)
+ (map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns);
+
+fun tyscm_rhss_of thy c eqns =
+ let
+ val tyscm = case eqns of [] => Code.default_typscheme thy c
+ | ((thm, _) :: _) => Code.typscheme_eqn thy thm;
+ val rhss = consts_of thy eqns;
+ in (tyscm, rhss) end;
+
+
+(* data structures *)
+
+datatype const = Fun of string | Inst of class * string;
+
+fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2)
+ | const_ord (Inst class_tyco1, Inst class_tyco2) =
+ prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2)
+ | const_ord (Fun _, Inst _) = LESS
+ | const_ord (Inst _, Fun _) = GREATER;
+
+type var = const * int;
+
+structure Vargraph =
+ GraphFun(type key = var val ord = prod_ord const_ord int_ord);
+
+datatype styp = Tyco of string * styp list | Var of var | Free;
+
+fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys)
+ | styp_of c_lhs (TFree (v, _)) = case c_lhs
+ of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs)
+ | NONE => Free;
+
+type vardeps_data = ((string * styp list) list * class list) Vargraph.T
+ * (((string * sort) list * (thm * bool) list) Symtab.table
+ * (class * string) list);
+
+val empty_vardeps_data : vardeps_data =
+ (Vargraph.empty, (Symtab.empty, []));
+
+
+(* retrieving equations and instances from the background context *)
+
+fun obtain_eqns thy eqngr c =
+ case try (Graph.get_node eqngr) c
+ of SOME ((lhs, _), eqns) => ((lhs, []), [])
+ | NONE => let
+ val eqns = Code.these_eqns thy c
+ |> preprocess thy c;
+ val ((lhs, _), rhss) = tyscm_rhss_of thy c eqns;
+ in ((lhs, rhss), eqns) end;
+
+fun obtain_instance thy arities (inst as (class, tyco)) =
+ case AList.lookup (op =) arities inst
+ of SOME classess => (classess, ([], []))
+ | NONE => let
+ val all_classes = complete_proper_sort thy [class];
+ val superclasses = remove (op =) class all_classes
+ val classess = map (complete_proper_sort thy)
+ (Sign.arity_sorts thy tyco [class]);
+ val inst_params = inst_params thy tyco all_classes;
+ in (classess, (superclasses, inst_params)) end;
+
+
+(* computing instantiations *)
+
+fun add_classes thy arities eqngr c_k new_classes vardeps_data =
+ let
+ val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ val diff_classes = new_classes |> subtract (op =) old_classes;
+ in if null diff_classes then vardeps_data
+ else let
+ val c_ks = Vargraph.imm_succs (fst vardeps_data) c_k |> insert (op =) c_k;
+ in
+ vardeps_data
+ |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes)
+ |> fold (fn styp => fold (ensure_typmatch_inst thy arities eqngr styp) new_classes) styps
+ |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
+ end end
+and add_styp thy arities eqngr c_k tyco_styps vardeps_data =
+ let
+ val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ in if member (op =) old_styps tyco_styps then vardeps_data
+ else
+ vardeps_data
+ |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps)
+ |> fold (ensure_typmatch_inst thy arities eqngr tyco_styps) classes
+ end
+and add_dep thy arities eqngr c_k c_k' vardeps_data =
+ let
+ val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ in
+ vardeps_data
+ |> add_classes thy arities eqngr c_k' classes
+ |> apfst (Vargraph.add_edge (c_k, c_k'))
+ end
+and ensure_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data =
+ if can (Sign.arity_sorts thy tyco) [class]
+ then vardeps_data
+ |> ensure_inst thy arities eqngr (class, tyco)
+ |> fold_index (fn (k, styp) =>
+ ensure_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps
+ else vardeps_data (*permissive!*)
+and ensure_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
+ if member (op =) insts inst then vardeps_data
+ else let
+ val (classess, (superclasses, inst_params)) =
+ obtain_instance thy arities inst;
+ in
+ vardeps_data
+ |> (apsnd o apsnd) (insert (op =) inst)
+ |> fold_index (fn (k, _) =>
+ apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
+ |> fold (fn superclass => ensure_inst thy arities eqngr (superclass, tyco)) superclasses
+ |> fold (ensure_fun thy arities eqngr) inst_params
+ |> fold_index (fn (k, classes) =>
+ add_classes thy arities eqngr (Inst (class, tyco), k) classes
+ #> fold (fn superclass =>
+ add_dep thy arities eqngr (Inst (superclass, tyco), k)
+ (Inst (class, tyco), k)) superclasses
+ #> fold (fn inst_param =>
+ add_dep thy arities eqngr (Fun inst_param, k)
+ (Inst (class, tyco), k)
+ ) inst_params
+ ) classess
+ end
+and ensure_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data =
+ vardeps_data
+ |> add_styp thy arities eqngr c_k tyco_styps
+ | ensure_typmatch thy arities eqngr (Var c_k') c_k vardeps_data =
+ vardeps_data
+ |> add_dep thy arities eqngr c_k c_k'
+ | ensure_typmatch thy arities eqngr Free c_k vardeps_data =
+ vardeps_data
+and ensure_rhs thy arities eqngr (c', styps) vardeps_data =
+ vardeps_data
+ |> ensure_fun thy arities eqngr c'
+ |> fold_index (fn (k, styp) =>
+ ensure_typmatch thy arities eqngr styp (Fun c', k)) styps
+and ensure_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) =
+ if Symtab.defined eqntab c then vardeps_data
+ else let
+ val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c;
+ val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss;
+ in
+ vardeps_data
+ |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns)))
+ |> fold_index (fn (k, _) =>
+ apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs
+ |> fold_index (fn (k, (_, sort)) =>
+ add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs
+ |> fold (ensure_rhs thy arities eqngr) rhss'
+ end;
+
+
+(* applying instantiations *)
+
+fun dicts_of thy (proj_sort, algebra) (T, sort) =
+ let
+ fun class_relation (x, _) _ = x;
+ fun type_constructor tyco xs class =
+ inst_params thy tyco (Sorts.complete_sort algebra [class])
+ @ (maps o maps) fst xs;
+ fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
+ in
+ flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra
+ { class_relation = class_relation, type_constructor = type_constructor,
+ type_variable = type_variable } (T, proj_sort sort)
+ handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
+ end;
+
+fun add_arity thy vardeps (class, tyco) =
+ AList.default (op =)
+ ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
+ (0 upto Sign.arity_number thy tyco - 1));
+
+fun add_eqs thy vardeps (c, (proto_lhs, proto_eqns)) (rhss, eqngr) =
+ if can (Graph.get_node eqngr) c then (rhss, eqngr)
+ else let
+ val lhs = map_index (fn (k, (v, _)) =>
+ (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs;
+ val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
+ Vartab.update ((v, 0), sort)) lhs;
+ val eqns = proto_eqns
+ |> (map o apfst) (Code.inst_thm thy inst_tab);
+ val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
+ val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
+ in (map (pair c) rhss' @ rhss, eqngr') end;
+
+fun extend_arities_eqngr thy cs ts (arities, eqngr) =
+ let
+ val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) =>
+ insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts [];
+ val (vardeps, (eqntab, insts)) = empty_vardeps_data
+ |> fold (ensure_fun thy arities eqngr) cs
+ |> fold (ensure_rhs thy arities eqngr) cs_rhss;
+ val arities' = fold (add_arity thy vardeps) insts arities;
+ val pp = Syntax.pp_global thy;
+ val algebra = Sorts.subalgebra pp (is_proper_class thy)
+ (AList.lookup (op =) arities') (Sign.classes_of thy);
+ val (rhss, eqngr') = Symtab.fold (add_eqs thy vardeps) eqntab ([], eqngr);
+ fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra)
+ (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c);
+ val eqngr'' = fold (fn (c, rhs) => fold
+ (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr';
+ in (algebra, (arities', eqngr'')) end;
+
+
+(** store for preprocessed arities and code equations **)
+
+structure Wellsorted = CodeDataFun
+(
+ type T = ((string * class) * sort list) list * code_graph;
+ val empty = ([], Graph.empty);
+ fun purge thy cs (arities, eqngr) =
+ let
+ val del_cs = ((Graph.all_preds eqngr
+ o filter (can (Graph.get_node eqngr))) cs);
+ val del_arities = del_cs
+ |> map_filter (AxClass.inst_of_param thy)
+ |> maps (fn (c, tyco) =>
+ (map (rpair tyco) o Sign.complete_sort thy o the_list
+ o AxClass.class_of_param thy) c);
+ val arities' = fold (AList.delete (op =)) del_arities arities;
+ val eqngr' = Graph.del_nodes del_cs eqngr;
+ in (arities', eqngr') end;
+);
+
+
+(** retrieval and evaluation interfaces **)
+
+fun obtain thy cs ts = apsnd snd
+ (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
+
+fun prepare_sorts_typ prep_sort
+ = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
+
+fun prepare_sorts prep_sort (Const (c, ty)) =
+ Const (c, prepare_sorts_typ prep_sort ty)
+ | prepare_sorts prep_sort (t1 $ t2) =
+ prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
+ | prepare_sorts prep_sort (Abs (v, ty, t)) =
+ Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
+ | prepare_sorts _ (t as Bound _) = t;
+
+fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct =
+ let
+ val pp = Syntax.pp_global thy;
+ val ct = cterm_of proto_ct;
+ val _ = (Sign.no_frees pp o map_types (K dummyT) o Sign.no_vars pp)
+ (Thm.term_of ct);
+ val thm = preprocess_conv thy ct;
+ val ct' = Thm.rhs_of thm;
+ val t' = Thm.term_of ct';
+ val vs = Term.add_tfrees t' [];
+ val consts = fold_aterms
+ (fn Const (c, _) => insert (op =) c | _ => I) t' [];
+
+ val t'' = prepare_sorts prep_sort t';
+ val (algebra', eqngr') = obtain thy consts [t''];
+ in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
+
+fun simple_evaluator evaluator algebra eqngr vs t ct =
+ evaluator algebra eqngr vs t;
+
+fun eval_conv thy =
+ let
+ fun conclude_evaluation thm2 thm1 =
+ let
+ val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
+ in
+ Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
+ error ("could not construct evaluation proof:\n"
+ ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3])
+ end;
+ in gen_eval thy I conclude_evaluation end;
+
+fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy)
+ (K o postproc (postprocess_term thy)) prep_sort (simple_evaluator evaluator);
+
+
+(** setup **)
+
+val setup =
+ let
+ fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+ fun add_del_attribute (name, (add, del)) =
+ Code.add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
+ || Scan.succeed (mk_attribute add))
+ in
+ add_del_attribute ("inline", (add_inline, del_inline))
+ #> add_del_attribute ("post", (add_post, del_post))
+ #> Code.add_attribute ("unfold", Scan.succeed (Thm.declaration_attribute
+ (fn thm => Context.mapping (Codegen.add_unfold thm #> add_inline thm) I)))
+ end;
+
+val _ =
+ OuterSyntax.improper_command "print_codeproc" "print code preprocessor setup"
+ OuterKeyword.diag (Scan.succeed
+ (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
+ (print_codeproc o Toplevel.theory_of)));
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_printer.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,318 @@
+(* Title: Tools/code/code_printer.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Generic operations for pretty printing of target language code.
+*)
+
+signature CODE_PRINTER =
+sig
+ val nerror: thm -> string -> 'a
+
+ val @@ : 'a * 'a -> 'a list
+ val @| : 'a list * 'a -> 'a list
+ val str: string -> Pretty.T
+ val concat: Pretty.T list -> Pretty.T
+ val brackets: Pretty.T list -> Pretty.T
+ val semicolon: Pretty.T list -> Pretty.T
+ val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T
+
+ val first_upper: string -> string
+ val first_lower: string -> string
+ type var_ctxt
+ val make_vars: string list -> var_ctxt
+ val intro_vars: string list -> var_ctxt -> var_ctxt
+ val lookup_var: var_ctxt -> string -> string
+
+ type literals
+ val Literals: { literal_char: string -> string, literal_string: string -> string,
+ literal_numeral: bool -> int -> string,
+ literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }
+ -> literals
+ val literal_char: literals -> string -> string
+ val literal_string: literals -> string -> string
+ val literal_numeral: literals -> bool -> int -> string
+ val literal_list: literals -> Pretty.T list -> Pretty.T
+ val infix_cons: literals -> int * string
+
+ type lrx
+ val L: lrx
+ val R: lrx
+ val X: lrx
+ type fixity
+ val BR: fixity
+ val NOBR: fixity
+ val INFX: int * lrx -> fixity
+ val APP: fixity
+ val brackify: fixity -> Pretty.T list -> Pretty.T
+ val brackify_infix: int * lrx -> fixity -> Pretty.T list -> Pretty.T
+ val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T
+
+ type itype = Code_Thingol.itype
+ type iterm = Code_Thingol.iterm
+ type const = Code_Thingol.const
+ type dict = Code_Thingol.dict
+ type tyco_syntax
+ type const_syntax
+ type proto_const_syntax
+ val parse_infix: ('a -> 'b) -> lrx * int -> string
+ -> int * ((fixity -> 'b -> Pretty.T)
+ -> fixity -> 'a list -> Pretty.T)
+ val parse_syntax: ('a -> 'b) -> OuterParse.token list
+ -> (int * ((fixity -> 'b -> Pretty.T)
+ -> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list
+ val simple_const_syntax: (int * ((fixity -> iterm -> Pretty.T)
+ -> fixity -> (iterm * itype) list -> Pretty.T)) option -> proto_const_syntax option
+ val activate_const_syntax: theory -> literals
+ -> proto_const_syntax -> Code_Thingol.naming -> const_syntax * Code_Thingol.naming
+ val gen_pr_app: (thm -> var_ctxt -> const * iterm list -> Pretty.T list)
+ -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
+ -> (string -> const_syntax option)
+ -> thm -> var_ctxt -> fixity -> const * iterm list -> Pretty.T
+ val gen_pr_bind: ((string option * Pretty.T option) * itype -> Pretty.T)
+ -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
+ -> thm -> fixity
+ -> (string option * iterm option) * itype -> var_ctxt -> Pretty.T * var_ctxt
+
+ val mk_name_module: Name.context -> string option -> (string -> string option)
+ -> 'a Graph.T -> string -> string
+ val dest_name: string -> string * string
+end;
+
+structure Code_Printer : CODE_PRINTER =
+struct
+
+open Code_Thingol;
+
+fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm thm);
+
+(** assembling text pieces **)
+
+infixr 5 @@;
+infixr 5 @|;
+fun x @@ y = [x, y];
+fun xs @| y = xs @ [y];
+val str = PrintMode.setmp [] Pretty.str;
+val concat = Pretty.block o Pretty.breaks;
+val brackets = Pretty.enclose "(" ")" o Pretty.breaks;
+fun semicolon ps = Pretty.block [concat ps, str ";"];
+fun enum_default default sep opn cls [] = str default
+ | enum_default default sep opn cls xs = Pretty.enum sep opn cls xs;
+
+
+(** names and variable name contexts **)
+
+type var_ctxt = string Symtab.table * Name.context;
+
+fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
+ Name.make_context names);
+
+fun intro_vars names (namemap, namectxt) =
+ let
+ val (names', namectxt') = Name.variants names namectxt;
+ val namemap' = fold2 (curry Symtab.update) names names' namemap;
+ in (namemap', namectxt') end;
+
+fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
+ of SOME name' => name'
+ | NONE => error ("Invalid name in context: " ^ quote name);
+
+val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
+val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
+
+
+(** pretty literals **)
+
+datatype literals = Literals of {
+ literal_char: string -> string,
+ literal_string: string -> string,
+ literal_numeral: bool -> int -> string,
+ literal_list: Pretty.T list -> Pretty.T,
+ infix_cons: int * string
+};
+
+fun dest_Literals (Literals lits) = lits;
+
+val literal_char = #literal_char o dest_Literals;
+val literal_string = #literal_string o dest_Literals;
+val literal_numeral = #literal_numeral o dest_Literals;
+val literal_list = #literal_list o dest_Literals;
+val infix_cons = #infix_cons o dest_Literals;
+
+
+(** syntax printer **)
+
+(* binding priorities *)
+
+datatype lrx = L | R | X;
+
+datatype fixity =
+ BR
+ | NOBR
+ | INFX of (int * lrx);
+
+val APP = INFX (~1, L);
+
+fun fixity_lrx L L = false
+ | fixity_lrx R R = false
+ | fixity_lrx _ _ = true;
+
+fun fixity NOBR _ = false
+ | fixity _ NOBR = false
+ | fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) =
+ pr < pr_ctxt
+ orelse pr = pr_ctxt
+ andalso fixity_lrx lr lr_ctxt
+ orelse pr_ctxt = ~1
+ | fixity BR (INFX _) = false
+ | fixity _ _ = true;
+
+fun gen_brackify _ [p] = p
+ | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
+ | gen_brackify false (ps as _::_) = Pretty.block ps;
+
+fun brackify fxy_ctxt =
+ gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks;
+
+fun brackify_infix infx fxy_ctxt =
+ gen_brackify (fixity (INFX infx) fxy_ctxt) o Pretty.breaks;
+
+fun brackify_block fxy_ctxt p1 ps p2 =
+ let val p = Pretty.block_enclose (p1, p2) ps
+ in if fixity BR fxy_ctxt
+ then Pretty.enclose "(" ")" [p]
+ else p
+ end;
+
+
+(* generic syntax *)
+
+type tyco_syntax = int * ((fixity -> itype -> Pretty.T)
+ -> fixity -> itype list -> Pretty.T);
+type const_syntax = int * ((var_ctxt -> fixity -> iterm -> Pretty.T)
+ -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
+type proto_const_syntax = int * (string list * (literals -> string list
+ -> (var_ctxt -> fixity -> iterm -> Pretty.T)
+ -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T));
+
+fun simple_const_syntax (SOME (n, f)) = SOME (n,
+ ([], (fn _ => fn _ => fn pr => fn thm => fn vars => f (pr vars))))
+ | simple_const_syntax NONE = NONE;
+
+fun activate_const_syntax thy literals (n, (cs, f)) naming =
+ fold_map (Code_Thingol.ensure_declared_const thy) cs naming
+ |-> (fn cs' => pair (n, f literals cs'));
+
+fun gen_pr_app pr_app pr_term syntax_const thm vars fxy (app as ((c, (_, tys)), ts)) =
+ case syntax_const c
+ of NONE => brackify fxy (pr_app thm vars app)
+ | SOME (k, pr) =>
+ let
+ fun pr' fxy ts = pr (pr_term thm) thm vars fxy (ts ~~ curry Library.take k tys);
+ in if k = length ts
+ then pr' fxy ts
+ else if k < length ts
+ then case chop k ts of (ts1, ts2) =>
+ brackify fxy (pr' APP ts1 :: map (pr_term thm vars BR) ts2)
+ else pr_term thm vars fxy (Code_Thingol.eta_expand k app)
+ end;
+
+fun gen_pr_bind pr_bind pr_term thm (fxy : fixity) ((v, pat), ty : itype) vars =
+ let
+ val vs = case pat
+ of SOME pat => Code_Thingol.fold_varnames (insert (op =)) pat []
+ | NONE => [];
+ val vars' = intro_vars (the_list v) vars;
+ val vars'' = intro_vars vs vars';
+ val v' = Option.map (lookup_var vars') v;
+ val pat' = Option.map (pr_term thm vars'' fxy) pat;
+ in (pr_bind ((v', pat'), ty), vars'') end;
+
+
+(* mixfix syntax *)
+
+datatype 'a mixfix =
+ Arg of fixity
+ | Pretty of Pretty.T;
+
+fun mk_mixfix prep_arg (fixity_this, mfx) =
+ let
+ fun is_arg (Arg _) = true
+ | is_arg _ = false;
+ val i = (length o filter is_arg) mfx;
+ fun fillin _ [] [] =
+ []
+ | fillin pr (Arg fxy :: mfx) (a :: args) =
+ (pr fxy o prep_arg) a :: fillin pr mfx args
+ | fillin pr (Pretty p :: mfx) args =
+ p :: fillin pr mfx args;
+ in
+ (i, fn pr => fn fixity_ctxt => fn args =>
+ gen_brackify (fixity fixity_this fixity_ctxt) (fillin pr mfx args))
+ end;
+
+fun parse_infix prep_arg (x, i) s =
+ let
+ val l = case x of L => INFX (i, L) | _ => INFX (i, X);
+ val r = case x of R => INFX (i, R) | _ => INFX (i, X);
+ in
+ mk_mixfix prep_arg (INFX (i, x),
+ [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r])
+ end;
+
+fun parse_mixfix prep_arg s =
+ let
+ val sym_any = Scan.one Symbol.is_regular;
+ val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat (
+ ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
+ || ($$ "_" >> K (Arg BR))
+ || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length))
+ || (Scan.repeat1
+ ( $$ "'" |-- sym_any
+ || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
+ sym_any) >> (Pretty o str o implode)));
+ in case Scan.finite Symbol.stopper parse (Symbol.explode s)
+ of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p)
+ | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p)
+ | _ => Scan.!!
+ (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail ()
+ end;
+
+val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr");
+
+fun parse_syntax prep_arg xs =
+ Scan.option ((
+ ((OuterParse.$$$ infixK >> K X)
+ || (OuterParse.$$$ infixlK >> K L)
+ || (OuterParse.$$$ infixrK >> K R))
+ -- OuterParse.nat >> parse_infix prep_arg
+ || Scan.succeed (parse_mixfix prep_arg))
+ -- OuterParse.string
+ >> (fn (parse, s) => parse s)) xs;
+
+val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK];
+
+
+(** module name spaces **)
+
+val dest_name =
+ apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
+
+fun mk_name_module reserved_names module_prefix module_alias program =
+ let
+ fun mk_alias name = case module_alias name
+ of SOME name' => name'
+ | NONE => name
+ |> Long_Name.explode
+ |> map (fn name => (the_single o fst) (Name.variants [name] reserved_names))
+ |> Long_Name.implode;
+ fun mk_prefix name = case module_prefix
+ of SOME module_prefix => Long_Name.append module_prefix name
+ | NONE => name;
+ val tab =
+ Symtab.empty
+ |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
+ o fst o dest_name o fst)
+ program
+ in the o Symtab.lookup tab end;
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_target.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,629 @@
+(* Title: Tools/code/code_target.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Serializer from intermediate language ("Thin-gol") to target languages.
+*)
+
+signature CODE_TARGET =
+sig
+ include CODE_PRINTER
+
+ type serializer
+ val add_target: string * (serializer * literals) -> theory -> theory
+ val extend_target: string *
+ (string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program))
+ -> theory -> theory
+ val assert_target: theory -> string -> string
+
+ type destination
+ type serialization
+ val parse_args: (OuterLex.token list -> 'a * OuterLex.token list)
+ -> OuterLex.token list -> 'a
+ val stmt_names_of_destination: destination -> string list
+ val code_of_pretty: Pretty.T -> string
+ val code_writeln: Pretty.T -> unit
+ val mk_serialization: string -> ('a -> unit) option
+ -> (Path.T option -> 'a -> unit)
+ -> ('a -> string * string option list)
+ -> 'a -> serialization
+ val serialize: theory -> string -> string option -> OuterLex.token list
+ -> Code_Thingol.naming -> Code_Thingol.program -> string list -> serialization
+ val serialize_custom: theory -> string * (serializer * literals)
+ -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
+ val the_literals: theory -> string -> literals
+ val compile: serialization -> unit
+ val export: serialization -> unit
+ val file: Path.T -> serialization -> unit
+ val string: string list -> serialization -> string
+ val code_of: theory -> string -> string
+ -> string list -> (Code_Thingol.naming -> string list) -> string
+ val shell_command: string (*theory name*) -> string (*export_code expr*) -> unit
+ val code_width: int ref
+
+ val allow_abort: string -> theory -> theory
+ val add_syntax_class: string -> class -> string option -> theory -> theory
+ val add_syntax_inst: string -> string * class -> bool -> theory -> theory
+ val add_syntax_tyco: string -> string -> tyco_syntax option -> theory -> theory
+ val add_syntax_const: string -> string -> proto_const_syntax option -> theory -> theory
+ val add_reserved: string -> string -> theory -> theory
+end;
+
+structure Code_Target : CODE_TARGET =
+struct
+
+open Basic_Code_Thingol;
+open Code_Printer;
+
+(** basics **)
+
+datatype destination = Compile | Export | File of Path.T | String of string list;
+type serialization = destination -> (string * string option list) option;
+
+val code_width = ref 80; (*FIXME after Pretty module no longer depends on print mode*)
+fun code_setmp f = PrintMode.setmp [] (Pretty.setmp_margin (!code_width) f);
+fun code_of_pretty p = code_setmp Pretty.string_of p ^ "\n";
+fun code_writeln p = Pretty.setmp_margin (!code_width) Pretty.writeln p;
+
+(*FIXME why another code_setmp?*)
+fun compile f = (code_setmp f Compile; ());
+fun export f = (code_setmp f Export; ());
+fun file p f = (code_setmp f (File p); ());
+fun string stmts f = fst (the (code_setmp f (String stmts)));
+
+fun stmt_names_of_destination (String stmts) = stmts
+ | stmt_names_of_destination _ = [];
+
+fun mk_serialization target (SOME comp) _ _ code Compile = (comp code; NONE)
+ | mk_serialization target NONE _ _ _ Compile = error (target ^ ": no internal compilation")
+ | mk_serialization target _ output _ code Export = (output NONE code ; NONE)
+ | mk_serialization target _ output _ code (File file) = (output (SOME file) code; NONE)
+ | mk_serialization target _ _ string code (String _) = SOME (string code);
+
+
+(** theory data **)
+
+datatype name_syntax_table = NameSyntaxTable of {
+ class: string Symtab.table,
+ instance: unit Symreltab.table,
+ tyco: tyco_syntax Symtab.table,
+ const: proto_const_syntax Symtab.table
+};
+
+fun mk_name_syntax_table ((class, instance), (tyco, const)) =
+ NameSyntaxTable { class = class, instance = instance, tyco = tyco, const = const };
+fun map_name_syntax_table f (NameSyntaxTable { class, instance, tyco, const }) =
+ mk_name_syntax_table (f ((class, instance), (tyco, const)));
+fun merge_name_syntax_table (NameSyntaxTable { class = class1, instance = instance1, tyco = tyco1, const = const1 },
+ NameSyntaxTable { class = class2, instance = instance2, tyco = tyco2, const = const2 }) =
+ mk_name_syntax_table (
+ (Symtab.join (K snd) (class1, class2),
+ Symreltab.join (K snd) (instance1, instance2)),
+ (Symtab.join (K snd) (tyco1, tyco2),
+ Symtab.join (K snd) (const1, const2))
+ );
+
+type serializer =
+ string option (*module name*)
+ -> OuterLex.token list (*arguments*)
+ -> (string -> string) (*labelled_name*)
+ -> string list (*reserved symbols*)
+ -> (string * Pretty.T) list (*includes*)
+ -> (string -> string option) (*module aliasses*)
+ -> (string -> string option) (*class syntax*)
+ -> (string -> tyco_syntax option)
+ -> (string -> const_syntax option)
+ -> Code_Thingol.program
+ -> string list (*selected statements*)
+ -> serialization;
+
+datatype serializer_entry = Serializer of serializer * literals
+ | Extends of string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program);
+
+datatype target = Target of {
+ serial: serial,
+ serializer: serializer_entry,
+ reserved: string list,
+ includes: (Pretty.T * string list) Symtab.table,
+ name_syntax_table: name_syntax_table,
+ module_alias: string Symtab.table
+};
+
+fun make_target ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))) =
+ Target { serial = serial, serializer = serializer, reserved = reserved,
+ includes = includes, name_syntax_table = name_syntax_table, module_alias = module_alias };
+fun map_target f ( Target { serial, serializer, reserved, includes, name_syntax_table, module_alias } ) =
+ make_target (f ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))));
+fun merge_target strict target (Target { serial = serial1, serializer = serializer,
+ reserved = reserved1, includes = includes1,
+ name_syntax_table = name_syntax_table1, module_alias = module_alias1 },
+ Target { serial = serial2, serializer = _,
+ reserved = reserved2, includes = includes2,
+ name_syntax_table = name_syntax_table2, module_alias = module_alias2 }) =
+ if serial1 = serial2 orelse not strict then
+ make_target ((serial1, serializer),
+ ((merge (op =) (reserved1, reserved2), Symtab.merge (op =) (includes1, includes2)),
+ (merge_name_syntax_table (name_syntax_table1, name_syntax_table2),
+ Symtab.join (K snd) (module_alias1, module_alias2))
+ ))
+ else
+ error ("Incompatible serializers: " ^ quote target);
+
+structure CodeTargetData = TheoryDataFun
+(
+ type T = target Symtab.table * string list;
+ val empty = (Symtab.empty, []);
+ val copy = I;
+ val extend = I;
+ fun merge _ ((target1, exc1) : T, (target2, exc2)) =
+ (Symtab.join (merge_target true) (target1, target2), Library.merge (op =) (exc1, exc2));
+);
+
+fun the_serializer (Target { serializer, ... }) = serializer;
+fun the_reserved (Target { reserved, ... }) = reserved;
+fun the_includes (Target { includes, ... }) = includes;
+fun the_name_syntax (Target { name_syntax_table = NameSyntaxTable x, ... }) = x;
+fun the_module_alias (Target { module_alias , ... }) = module_alias;
+
+val abort_allowed = snd o CodeTargetData.get;
+
+fun assert_target thy target =
+ case Symtab.lookup (fst (CodeTargetData.get thy)) target
+ of SOME data => target
+ | NONE => error ("Unknown code target language: " ^ quote target);
+
+fun put_target (target, seri) thy =
+ let
+ val lookup_target = Symtab.lookup (fst (CodeTargetData.get thy));
+ val _ = case seri
+ of Extends (super, _) => if is_some (lookup_target super) then ()
+ else error ("Unknown code target language: " ^ quote super)
+ | _ => ();
+ val overwriting = case (Option.map the_serializer o lookup_target) target
+ of NONE => false
+ | SOME (Extends _) => true
+ | SOME (Serializer _) => (case seri
+ of Extends _ => error ("Will not overwrite existing target " ^ quote target)
+ | _ => true);
+ val _ = if overwriting
+ then warning ("Overwriting existing target " ^ quote target)
+ else ();
+ in
+ thy
+ |> (CodeTargetData.map o apfst oo Symtab.map_default)
+ (target, make_target ((serial (), seri), (([], Symtab.empty),
+ (mk_name_syntax_table ((Symtab.empty, Symreltab.empty), (Symtab.empty, Symtab.empty)),
+ Symtab.empty))))
+ ((map_target o apfst o apsnd o K) seri)
+ end;
+
+fun add_target (target, seri) = put_target (target, Serializer seri);
+fun extend_target (target, (super, modify)) =
+ put_target (target, Extends (super, modify));
+
+fun map_target_data target f thy =
+ let
+ val _ = assert_target thy target;
+ in
+ thy
+ |> (CodeTargetData.map o apfst o Symtab.map_entry target o map_target) f
+ end;
+
+fun map_reserved target =
+ map_target_data target o apsnd o apfst o apfst;
+fun map_includes target =
+ map_target_data target o apsnd o apfst o apsnd;
+fun map_name_syntax target =
+ map_target_data target o apsnd o apsnd o apfst o map_name_syntax_table;
+fun map_module_alias target =
+ map_target_data target o apsnd o apsnd o apsnd;
+
+
+(** serializer configuration **)
+
+(* data access *)
+
+local
+
+fun cert_class thy class =
+ let
+ val _ = AxClass.get_info thy class;
+ in class end;
+
+fun read_class thy = cert_class thy o Sign.intern_class thy;
+
+fun cert_tyco thy tyco =
+ let
+ val _ = if Sign.declared_tyname thy tyco then ()
+ else error ("No such type constructor: " ^ quote tyco);
+ in tyco end;
+
+fun read_tyco thy = cert_tyco thy o Sign.intern_type thy;
+
+fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy =
+ let
+ val class = prep_class thy raw_class;
+ in case raw_syn
+ of SOME syntax =>
+ thy
+ |> (map_name_syntax target o apfst o apfst)
+ (Symtab.update (class, syntax))
+ | NONE =>
+ thy
+ |> (map_name_syntax target o apfst o apfst)
+ (Symtab.delete_safe class)
+ end;
+
+fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy =
+ let
+ val inst = (prep_class thy raw_class, prep_tyco thy raw_tyco);
+ in if add_del then
+ thy
+ |> (map_name_syntax target o apfst o apsnd)
+ (Symreltab.update (inst, ()))
+ else
+ thy
+ |> (map_name_syntax target o apfst o apsnd)
+ (Symreltab.delete_safe inst)
+ end;
+
+fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy =
+ let
+ val tyco = prep_tyco thy raw_tyco;
+ fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco
+ then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
+ else syntax
+ in case raw_syn
+ of SOME syntax =>
+ thy
+ |> (map_name_syntax target o apsnd o apfst)
+ (Symtab.update (tyco, check_args syntax))
+ | NONE =>
+ thy
+ |> (map_name_syntax target o apsnd o apfst)
+ (Symtab.delete_safe tyco)
+ end;
+
+fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
+ let
+ val c = prep_const thy raw_c;
+ fun check_args (syntax as (n, _)) = if n > Code.no_args thy c
+ then error ("Too many arguments in syntax for constant " ^ quote c)
+ else syntax;
+ in case raw_syn
+ of SOME syntax =>
+ thy
+ |> (map_name_syntax target o apsnd o apsnd)
+ (Symtab.update (c, check_args syntax))
+ | NONE =>
+ thy
+ |> (map_name_syntax target o apsnd o apsnd)
+ (Symtab.delete_safe c)
+ end;
+
+fun add_reserved target =
+ let
+ fun add sym syms = if member (op =) syms sym
+ then error ("Reserved symbol " ^ quote sym ^ " already declared")
+ else insert (op =) sym syms
+ in map_reserved target o add end;
+
+fun gen_add_include read_const target args thy =
+ let
+ fun add (name, SOME (content, raw_cs)) incls =
+ let
+ val _ = if Symtab.defined incls name
+ then warning ("Overwriting existing include " ^ name)
+ else ();
+ val cs = map (read_const thy) raw_cs;
+ in Symtab.update (name, (str content, cs)) incls end
+ | add (name, NONE) incls = Symtab.delete name incls;
+ in map_includes target (add args) thy end;
+
+val add_include = gen_add_include Code.check_const;
+val add_include_cmd = gen_add_include Code.read_const;
+
+fun add_module_alias target (thyname, modlname) =
+ let
+ val xs = Long_Name.explode modlname;
+ val xs' = map (Name.desymbolize true) xs;
+ in if xs' = xs
+ then map_module_alias target (Symtab.update (thyname, modlname))
+ else error ("Invalid module name: " ^ quote modlname ^ "\n"
+ ^ "perhaps try " ^ quote (Long_Name.implode xs'))
+ end;
+
+fun gen_allow_abort prep_const raw_c thy =
+ let
+ val c = prep_const thy raw_c;
+ in thy |> (CodeTargetData.map o apsnd) (insert (op =) c) end;
+
+fun zip_list (x::xs) f g =
+ f
+ #-> (fn y =>
+ fold_map (fn x => g |-- f >> pair x) xs
+ #-> (fn xys => pair ((x, y) :: xys)));
+
+
+(* concrete syntax *)
+
+structure P = OuterParse
+and K = OuterKeyword
+
+fun parse_multi_syntax parse_thing parse_syntax =
+ P.and_list1 parse_thing
+ #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
+ (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")"));
+
+in
+
+val add_syntax_class = gen_add_syntax_class cert_class (K I);
+val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco;
+val add_syntax_tyco = gen_add_syntax_tyco cert_tyco;
+val add_syntax_const = gen_add_syntax_const (K I);
+val allow_abort = gen_allow_abort (K I);
+val add_reserved = add_reserved;
+
+val add_syntax_class_cmd = gen_add_syntax_class read_class Code.read_const;
+val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
+val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
+val add_syntax_const_cmd = gen_add_syntax_const Code.read_const;
+val allow_abort_cmd = gen_allow_abort Code.read_const;
+
+fun the_literals thy =
+ let
+ val (targets, _) = CodeTargetData.get thy;
+ fun literals target = case Symtab.lookup targets target
+ of SOME data => (case the_serializer data
+ of Serializer (_, literals) => literals
+ | Extends (super, _) => literals super)
+ | NONE => error ("Unknown code target language: " ^ quote target);
+ in literals end;
+
+
+(** serializer usage **)
+
+(* montage *)
+
+local
+
+fun labelled_name thy program name = case Graph.get_node program name
+ of Code_Thingol.Fun (c, _) => quote (Code.string_of_const thy c)
+ | Code_Thingol.Datatype (tyco, _) => "type " ^ quote (Sign.extern_type thy tyco)
+ | Code_Thingol.Datatypecons (c, _) => quote (Code.string_of_const thy c)
+ | Code_Thingol.Class (class, _) => "class " ^ quote (Sign.extern_class thy class)
+ | Code_Thingol.Classrel (sub, super) => let
+ val Code_Thingol.Class (sub, _) = Graph.get_node program sub
+ val Code_Thingol.Class (super, _) = Graph.get_node program super
+ in quote (Sign.extern_class thy sub ^ " < " ^ Sign.extern_class thy super) end
+ | Code_Thingol.Classparam (c, _) => quote (Code.string_of_const thy c)
+ | Code_Thingol.Classinst ((class, (tyco, _)), _) => let
+ val Code_Thingol.Class (class, _) = Graph.get_node program class
+ val Code_Thingol.Datatype (tyco, _) = Graph.get_node program tyco
+ in quote (Sign.extern_type thy tyco ^ " :: " ^ Sign.extern_class thy class) end
+
+fun activate_syntax lookup_name src_tab = Symtab.empty
+ |> fold_map (fn thing_identifier => fn tab => case lookup_name thing_identifier
+ of SOME name => (SOME name,
+ Symtab.update_new (name, the (Symtab.lookup src_tab thing_identifier)) tab)
+ | NONE => (NONE, tab)) (Symtab.keys src_tab)
+ |>> map_filter I;
+
+fun activate_const_syntax thy literals src_tab naming = (Symtab.empty, naming)
+ |> fold_map (fn thing_identifier => fn (tab, naming) =>
+ case Code_Thingol.lookup_const naming thing_identifier
+ of SOME name => let
+ val (syn, naming') = Code_Printer.activate_const_syntax thy
+ literals (the (Symtab.lookup src_tab thing_identifier)) naming
+ in (SOME name, (Symtab.update_new (name, syn) tab, naming')) end
+ | NONE => (NONE, (tab, naming))) (Symtab.keys src_tab)
+ |>> map_filter I;
+
+fun invoke_serializer thy abortable serializer literals reserved abs_includes
+ module_alias class instance tyco const module args naming program2 names1 =
+ let
+ val (names_class, class') =
+ activate_syntax (Code_Thingol.lookup_class naming) class;
+ val names_inst = map_filter (Code_Thingol.lookup_instance naming)
+ (Symreltab.keys instance);
+ val (names_tyco, tyco') =
+ activate_syntax (Code_Thingol.lookup_tyco naming) tyco;
+ val (names_const, (const', _)) =
+ activate_const_syntax thy literals const naming;
+ val names_hidden = names_class @ names_inst @ names_tyco @ names_const;
+ val names2 = subtract (op =) names_hidden names1;
+ val program3 = Graph.subgraph (not o member (op =) names_hidden) program2;
+ val names_all = Graph.all_succs program3 names2;
+ val includes = abs_includes names_all;
+ val program4 = Graph.subgraph (member (op =) names_all) program3;
+ val empty_funs = filter_out (member (op =) abortable)
+ (Code_Thingol.empty_funs program3);
+ val _ = if null empty_funs then () else error ("No code equations for "
+ ^ commas (map (Sign.extern_const thy) empty_funs));
+ in
+ serializer module args (labelled_name thy program2) reserved includes
+ (Symtab.lookup module_alias) (Symtab.lookup class')
+ (Symtab.lookup tyco') (Symtab.lookup const')
+ program4 names2
+ end;
+
+fun mount_serializer thy alt_serializer target module args naming program names =
+ let
+ val (targets, abortable) = CodeTargetData.get thy;
+ fun collapse_hierarchy target =
+ let
+ val data = case Symtab.lookup targets target
+ of SOME data => data
+ | NONE => error ("Unknown code target language: " ^ quote target);
+ in case the_serializer data
+ of Serializer _ => (I, data)
+ | Extends (super, modify) => let
+ val (modify', data') = collapse_hierarchy super
+ in (modify' #> modify naming, merge_target false target (data', data)) end
+ end;
+ val (modify, data) = collapse_hierarchy target;
+ val (serializer, _) = the_default (case the_serializer data
+ of Serializer seri => seri) alt_serializer;
+ val reserved = the_reserved data;
+ fun select_include names_all (name, (content, cs)) =
+ if null cs then SOME (name, content)
+ else if exists (fn c => case Code_Thingol.lookup_const naming c
+ of SOME name => member (op =) names_all name
+ | NONE => false) cs
+ then SOME (name, content) else NONE;
+ fun includes names_all = map_filter (select_include names_all)
+ ((Symtab.dest o the_includes) data);
+ val module_alias = the_module_alias data;
+ val { class, instance, tyco, const } = the_name_syntax data;
+ val literals = the_literals thy target;
+ in
+ invoke_serializer thy abortable serializer literals reserved
+ includes module_alias class instance tyco const module args naming (modify program) names
+ end;
+
+in
+
+fun serialize thy = mount_serializer thy NONE;
+
+fun serialize_custom thy (target_name, seri) naming program names =
+ mount_serializer thy (SOME seri) target_name NONE [] naming program names (String [])
+ |> the;
+
+end; (* local *)
+
+fun parse_args f args =
+ case Scan.read OuterLex.stopper f args
+ of SOME x => x
+ | NONE => error "Bad serializer arguments";
+
+
+(* code presentation *)
+
+fun code_of thy target module_name cs names_stmt =
+ let
+ val (names_cs, (naming, program)) = Code_Thingol.consts_program thy cs;
+ in
+ string (names_stmt naming) (serialize thy target (SOME module_name) []
+ naming program names_cs)
+ end;
+
+
+(* code generation *)
+
+fun transitivly_non_empty_funs thy naming program =
+ let
+ val cs = subtract (op =) (abort_allowed thy) (Code_Thingol.empty_funs program);
+ val names = map_filter (Code_Thingol.lookup_const naming) cs;
+ in subtract (op =) (Graph.all_preds program names) (Graph.keys program) end;
+
+fun read_const_exprs thy cs =
+ let
+ val (cs1, cs2) = Code_Thingol.read_const_exprs thy cs;
+ val (names3, (naming, program)) = Code_Thingol.consts_program thy cs2;
+ val names4 = transitivly_non_empty_funs thy naming program;
+ val cs5 = map_filter
+ (fn (c, name) => if member (op =) names4 name then SOME c else NONE) (cs2 ~~ names3);
+ in fold (insert (op =)) cs5 cs1 end;
+
+fun cached_program thy =
+ let
+ val (naming, program) = Code_Thingol.cached_program thy;
+ in (transitivly_non_empty_funs thy naming program, (naming, program)) end
+
+fun export_code thy cs seris =
+ let
+ val (cs', (naming, program)) = if null cs then cached_program thy
+ else Code_Thingol.consts_program thy cs;
+ fun mk_seri_dest dest = case dest
+ of NONE => compile
+ | SOME "-" => export
+ | SOME f => file (Path.explode f)
+ val _ = map (fn (((target, module), dest), args) =>
+ (mk_seri_dest dest (serialize thy target module args naming program cs'))) seris;
+ in () end;
+
+fun export_code_cmd raw_cs seris thy = export_code thy (read_const_exprs thy raw_cs) seris;
+
+
+(** Isar setup **)
+
+val (inK, module_nameK, fileK) = ("in", "module_name", "file");
+
+val code_exprP =
+ (Scan.repeat P.term_group
+ -- Scan.repeat (P.$$$ inK |-- P.name
+ -- Scan.option (P.$$$ module_nameK |-- P.name)
+ -- Scan.option (P.$$$ fileK |-- P.name)
+ -- Scan.optional (P.$$$ "(" |-- Args.parse --| P.$$$ ")") []
+ ) >> (fn (raw_cs, seris) => export_code_cmd raw_cs seris));
+
+val _ = List.app OuterKeyword.keyword [inK, module_nameK, fileK];
+
+val _ =
+ OuterSyntax.command "code_class" "define code syntax for class" K.thy_decl (
+ parse_multi_syntax P.xname (Scan.option P.string)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns)
+ );
+
+val _ =
+ OuterSyntax.command "code_instance" "define code syntax for instance" K.thy_decl (
+ parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname)
+ ((P.minus >> K true) || Scan.succeed false)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns)
+ );
+
+val _ =
+ OuterSyntax.command "code_type" "define code syntax for type constructor" K.thy_decl (
+ parse_multi_syntax P.xname (parse_syntax I)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns)
+ );
+
+val _ =
+ OuterSyntax.command "code_const" "define code syntax for constant" K.thy_decl (
+ parse_multi_syntax P.term_group (parse_syntax fst)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_const, syn) => add_syntax_const_cmd target raw_const
+ (Code_Printer.simple_const_syntax syn)) syns)
+ );
+
+val _ =
+ OuterSyntax.command "code_reserved" "declare words as reserved for target language" K.thy_decl (
+ P.name -- Scan.repeat1 P.name
+ >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)
+ );
+
+val _ =
+ OuterSyntax.command "code_include" "declare piece of code to be included in generated code" K.thy_decl (
+ P.name -- P.name -- (P.text :|-- (fn "-" => Scan.succeed NONE
+ | s => Scan.optional (P.$$$ "attach" |-- Scan.repeat1 P.term) [] >> pair s >> SOME))
+ >> (fn ((target, name), content_consts) =>
+ (Toplevel.theory o add_include_cmd target) (name, content_consts))
+ );
+
+val _ =
+ OuterSyntax.command "code_modulename" "alias module to other name" K.thy_decl (
+ P.name -- Scan.repeat1 (P.name -- P.name)
+ >> (fn (target, modlnames) => (Toplevel.theory o fold (add_module_alias target)) modlnames)
+ );
+
+val _ =
+ OuterSyntax.command "code_abort" "permit constant to be implemented as program abort" K.thy_decl (
+ Scan.repeat1 P.term_group >> (Toplevel.theory o fold allow_abort_cmd)
+ );
+
+val _ =
+ OuterSyntax.command "export_code" "generate executable code for constants"
+ K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
+
+fun shell_command thyname cmd = Toplevel.program (fn _ =>
+ (use_thy thyname; case Scan.read OuterLex.stopper (P.!!! code_exprP)
+ ((filter OuterLex.is_proper o OuterSyntax.scan Position.none) cmd)
+ of SOME f => (writeln "Now generating code..."; f (theory thyname))
+ | NONE => error ("Bad directive " ^ quote cmd)))
+ handle TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure;
+
+end; (*local*)
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_thingol.ML Tue Jun 23 15:48:55 2009 +0100
@@ -0,0 +1,876 @@
+(* Title: Tools/code/code_thingol.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Intermediate language ("Thin-gol") representing executable code.
+Representation and translation.
+*)
+
+infix 8 `%%;
+infix 4 `$;
+infix 4 `$$;
+infixr 3 `|=>;
+infixr 3 `|==>;
+
+signature BASIC_CODE_THINGOL =
+sig
+ type vname = string;
+ datatype dict =
+ DictConst of string * dict list list
+ | DictVar of string list * (vname * (int * int));
+ datatype itype =
+ `%% of string * itype list
+ | ITyVar of vname;
+ type const = string * ((itype list * dict list list) * itype list (*types of arguments*))
+ datatype iterm =
+ IConst of const
+ | IVar of vname
+ | `$ of iterm * iterm
+ | `|=> of (vname * itype) * iterm
+ | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
+ (*((term, type), [(selector pattern, body term )]), primitive term)*)
+ val `$$ : iterm * iterm list -> iterm;
+ val `|==> : (vname * itype) list * iterm -> iterm;
+ type typscheme = (vname * sort) list * itype;
+end;
+
+signature CODE_THINGOL =
+sig
+ include BASIC_CODE_THINGOL
+ val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list
+ val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
+ val unfold_fun: itype -> itype list * itype
+ val unfold_app: iterm -> iterm * iterm list
+ val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option
+ val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm
+ val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option
+ val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm
+ val unfold_const_app: iterm -> (const * iterm list) option
+ val collapse_let: ((vname * itype) * iterm) * iterm
+ -> (iterm * itype) * (iterm * iterm) list
+ val eta_expand: int -> const * iterm list -> iterm
+ val contains_dictvar: iterm -> bool
+ val locally_monomorphic: iterm -> bool
+ val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
+ val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
+ val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
+
+ type naming
+ val empty_naming: naming
+ val lookup_class: naming -> class -> string option
+ val lookup_classrel: naming -> class * class -> string option
+ val lookup_tyco: naming -> string -> string option
+ val lookup_instance: naming -> class * string -> string option
+ val lookup_const: naming -> string -> string option
+ val ensure_declared_const: theory -> string -> naming -> string * naming
+
+ datatype stmt =
+ NoStmt
+ | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)
+ | Datatype of string * ((vname * sort) list * (string * itype list) list)
+ | Datatypecons of string * string
+ | Class of class * (vname * ((class * string) list * (string * itype) list))
+ | Classrel of class * class
+ | Classparam of string * class
+ | Classinst of (class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * ((string * const) * (thm * bool)) list)
+ type program = stmt Graph.T
+ val empty_funs: program -> string list
+ val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm
+ val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt
+ val is_cons: program -> string -> bool
+ val contr_classparam_typs: program -> string -> itype option list
+
+ val read_const_exprs: theory -> string list -> string list * string list
+ val consts_program: theory -> string list -> string list * (naming * program)
+ val cached_program: theory -> naming * program
+ val eval_conv: theory -> (sort -> sort)
+ -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm)
+ -> cterm -> thm
+ val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
+ -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
+ -> term -> 'a
+end;
+
+structure Code_Thingol: CODE_THINGOL =
+struct
+
+(** auxiliary **)
+
+fun unfoldl dest x =
+ case dest x
+ of NONE => (x, [])
+ | SOME (x1, x2) =>
+ let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end;
+
+fun unfoldr dest x =
+ case dest x
+ of NONE => ([], x)
+ | SOME (x1, x2) =>
+ let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end;
+
+
+(** language core - types, terms **)
+
+type vname = string;
+
+datatype dict =
+ DictConst of string * dict list list
+ | DictVar of string list * (vname * (int * int));
+
+datatype itype =
+ `%% of string * itype list
+ | ITyVar of vname;
+
+type const = string * ((itype list * dict list list) * itype list (*types of arguments*))
+
+datatype iterm =
+ IConst of const
+ | IVar of vname
+ | `$ of iterm * iterm
+ | `|=> of (vname * itype) * iterm
+ | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
+ (*see also signature*)
+
+val op `$$ = Library.foldl (op `$);
+val op `|==> = Library.foldr (op `|=>);
+
+val unfold_app = unfoldl
+ (fn op `$ t => SOME t
+ | _ => NONE);
+
+val split_abs =
+ (fn (v, ty) `|=> (t as ICase (((IVar w, _), [(p, t')]), _)) =>
+ if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t)
+ | (v, ty) `|=> t => SOME (((v, NONE), ty), t)
+ | _ => NONE);
+
+val unfold_abs = unfoldr split_abs;
+
+val split_let =
+ (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)
+ | _ => NONE);
+
+val unfold_let = unfoldr split_let;
+
+fun unfold_const_app t =
+ case unfold_app t
+ of (IConst c, ts) => SOME (c, ts)
+ | _ => NONE;
+
+fun fold_aiterms f (t as IConst _) = f t
+ | fold_aiterms f (t as IVar _) = f t
+ | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2
+ | fold_aiterms f (t as _ `|=> t') = f t #> fold_aiterms f t'
+ | fold_aiterms f (ICase (_, t)) = fold_aiterms f t;
+
+fun fold_constnames f =
+ let
+ fun add (IConst (c, _)) = f c
+ | add _ = I;
+ in fold_aiterms add end;
+
+fun fold_varnames f =
+ let
+ fun add (IVar v) = f v
+ | add ((v, _) `|=> _) = f v
+ | add _ = I;
+ in fold_aiterms add end;
+
+fun fold_unbound_varnames f =
+ let
+ fun add _ (IConst _) = I
+ | add vs (IVar v) = if not (member (op =) vs v) then f v else I
+ | add vs (t1 `$ t2) = add vs t1 #> add vs t2
+ | add vs ((v, _) `|=> t) = add (insert (op =) v vs) t
+ | add vs (ICase (_, t)) = add vs t;
+ in add [] end;
+
+fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) =
+ let
+ fun exists_v t = fold_unbound_varnames (fn w => fn b =>
+ b orelse v = w) t false;
+ in if v = w andalso forall (fn (t1, t2) =>
+ exists_v t1 orelse not (exists_v t2)) ds
+ then ((se, ty), ds)
+ else ((se, ty), [(IVar v, be)])
+ end
+ | collapse_let (((v, ty), se), be) =
+ ((se, ty), [(IVar v, be)])
+
+fun eta_expand k (c as (_, (_, tys)), ts) =
+ let
+ val j = length ts;
+ val l = k - j;
+ val ctxt = (fold o fold_varnames) Name.declare ts Name.context;
+ val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys);
+ in vs_tys `|==> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end;
+
+fun contains_dictvar t =
+ let
+ fun contains (DictConst (_, dss)) = (fold o fold) contains dss
+ | contains (DictVar _) = K true;
+ in
+ fold_aiterms
+ (fn IConst (_, ((_, dss), _)) => (fold o fold) contains dss | _ => I) t false
+ end;
+
+fun locally_monomorphic (IConst _) = false
+ | locally_monomorphic (IVar _) = true
+ | locally_monomorphic (t `$ _) = locally_monomorphic t
+ | locally_monomorphic (_ `|=> t) = locally_monomorphic t
+ | locally_monomorphic (ICase ((_, ds), _)) = exists (locally_monomorphic o snd) ds;
+
+
+(** namings **)
+
+(* policies *)
+
+local
+ fun thyname_of thy f x = the (AList.lookup (op =) (f x) Markup.theory_nameN);
+ fun thyname_of_class thy =
+ thyname_of thy (ProofContext.query_class (ProofContext.init thy));
+ fun thyname_of_tyco thy =
+ thyname_of thy (Type.the_tags (Sign.tsig_of thy));
+ fun thyname_of_instance thy inst = case AxClass.arity_property thy inst Markup.theory_nameN
+ of [] => error ("no such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
+ | thyname :: _ => thyname;
+ fun thyname_of_const thy c = case AxClass.class_of_param thy c
+ of SOME class => thyname_of_class thy class
+ | NONE => (case Code.get_datatype_of_constr thy c
+ of SOME dtco => thyname_of_tyco thy dtco
+ | NONE => thyname_of thy (Consts.the_tags (Sign.consts_of thy)) c);
+ fun purify_base "op &" = "and"
+ | purify_base "op |" = "or"
+ | purify_base "op -->" = "implies"
+ | purify_base "op :" = "member"
+ | purify_base "op =" = "eq"
+ | purify_base "*" = "product"
+ | purify_base "+" = "sum"
+ | purify_base s = Name.desymbolize false s;
+ fun namify thy get_basename get_thyname name =
+ let
+ val prefix = get_thyname thy name;
+ val base = (purify_base o get_basename) name;
+ in Long_Name.append prefix base end;
+in
+
+fun namify_class thy = namify thy Long_Name.base_name thyname_of_class;
+fun namify_classrel thy = namify thy (fn (class1, class2) =>
+ Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1) (fn thy => thyname_of_class thy o fst);
+ (*order fits nicely with composed projections*)
+fun namify_tyco thy "fun" = "Pure.fun"
+ | namify_tyco thy tyco = namify thy Long_Name.base_name thyname_of_tyco tyco;
+fun namify_instance thy = namify thy (fn (class, tyco) =>
+ Long_Name.base_name class ^ "_" ^ Long_Name.base_name tyco) thyname_of_instance;
+fun namify_const thy = namify thy Long_Name.base_name thyname_of_const;
+
+end; (* local *)
+
+
+(* data *)
+
+datatype naming = Naming of {
+ class: class Symtab.table * Name.context,
+ classrel: string Symreltab.table * Name.context,
+ tyco: string Symtab.table * Name.context,
+ instance: string Symreltab.table * Name.context,
+ const: string Symtab.table * Name.context
+}
+
+fun dest_Naming (Naming naming) = naming;
+
+val empty_naming = Naming {
+ class = (Symtab.empty, Name.context),
+ classrel = (Symreltab.empty, Name.context),
+ tyco = (Symtab.empty, Name.context),
+ instance = (Symreltab.empty, Name.context),
+ const = (Symtab.empty, Name.context)
+};
+
+local
+ fun mk_naming (class, classrel, tyco, instance, const) =
+ Naming { class = class, classrel = classrel,
+ tyco = tyco, instance = instance, const = const };
+ fun map_naming f (Naming { class, classrel, tyco, instance, const }) =
+ mk_naming (f (class, classrel, tyco, instance, const));
+in
+ fun map_class f = map_naming
+ (fn (class, classrel, tyco, inst, const) =>
+ (f class, classrel, tyco, inst, const));
+ fun map_classrel f = map_naming
+ (fn (class, classrel, tyco, inst, const) =>
+ (class, f classrel, tyco, inst, const));
+ fun map_tyco f = map_naming
+ (fn (class, classrel, tyco, inst, const) =>
+ (class, classrel, f tyco, inst, const));
+ fun map_instance f = map_naming
+ (fn (class, classrel, tyco, inst, const) =>
+ (class, classrel, tyco, f inst, const));
+ fun map_const f = map_naming
+ (fn (class, classrel, tyco, inst, const) =>
+ (class, classrel, tyco, inst, f const));
+end; (*local*)
+
+fun add_variant update (thing, name) (tab, used) =
+ let
+ val (name', used') = yield_singleton Name.variants name used;
+ val tab' = update (thing, name') tab;
+ in (tab', used') end;
+
+fun declare thy mapp lookup update namify thing =
+ mapp (add_variant update (thing, namify thy thing))
+ #> `(fn naming => the (lookup naming thing));
+
+
+(* lookup and declare *)
+
+local
+
+val suffix_class = "class";
+val suffix_classrel = "classrel"
+val suffix_tyco = "tyco";
+val suffix_instance = "inst";
+val suffix_const = "const";
+
+fun add_suffix nsp NONE = NONE
+ | add_suffix nsp (SOME name) = SOME (Long_Name.append name nsp);
+
+in
+
+val lookup_class = add_suffix suffix_class
+ oo Symtab.lookup o fst o #class o dest_Naming;
+val lookup_classrel = add_suffix suffix_classrel
+ oo Symreltab.lookup o fst o #classrel o dest_Naming;
+val lookup_tyco = add_suffix suffix_tyco
+ oo Symtab.lookup o fst o #tyco o dest_Naming;
+val lookup_instance = add_suffix suffix_instance
+ oo Symreltab.lookup o fst o #instance o dest_Naming;
+val lookup_const = add_suffix suffix_const
+ oo Symtab.lookup o fst o #const o dest_Naming;
+
+fun declare_class thy = declare thy map_class
+ lookup_class Symtab.update_new namify_class;
+fun declare_classrel thy = declare thy map_classrel
+ lookup_classrel Symreltab.update_new namify_classrel;
+fun declare_tyco thy = declare thy map_tyco
+ lookup_tyco Symtab.update_new namify_tyco;
+fun declare_instance thy = declare thy map_instance
+ lookup_instance Symreltab.update_new namify_instance;
+fun declare_const thy = declare thy map_const
+ lookup_const Symtab.update_new namify_const;
+
+fun ensure_declared_const thy const naming =
+ case lookup_const naming const
+ of SOME const' => (const', naming)
+ | NONE => declare_const thy const naming;
+
+val unfold_fun = unfoldr
+ (fn "Pure.fun.tyco" `%% [ty1, ty2] => SOME (ty1, ty2)
+ | _ => NONE); (*depends on suffix_tyco and namify_tyco!*)
+
+end; (* local *)
+
+
+(** statements, abstract programs **)
+
+type typscheme = (vname * sort) list * itype;
+datatype stmt =
+ NoStmt
+ | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)
+ | Datatype of string * ((vname * sort) list * (string * itype list) list)
+ | Datatypecons of string * string
+ | Class of class * (vname * ((class * string) list * (string * itype) list))
+ | Classrel of class * class
+ | Classparam of string * class
+ | Classinst of (class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * ((string * const) * (thm * bool)) list);
+
+type program = stmt Graph.T;
+
+fun empty_funs program =
+ Graph.fold (fn (name, (Fun (c, (_, [])), _)) => cons c
+ | _ => I) program [];
+
+fun map_terms_bottom_up f (t as IConst _) = f t
+ | map_terms_bottom_up f (t as IVar _) = f t
+ | map_terms_bottom_up f (t1 `$ t2) = f
+ (map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2)
+ | map_terms_bottom_up f ((v, ty) `|=> t) = f
+ ((v, ty) `|=> map_terms_bottom_up f t)
+ | map_terms_bottom_up f (ICase (((t, ty), ps), t0)) = f
+ (ICase (((map_terms_bottom_up f t, ty), (map o pairself)
+ (map_terms_bottom_up f) ps), map_terms_bottom_up f t0));
+
+fun map_terms_stmt f NoStmt = NoStmt
+ | map_terms_stmt f (Fun (c, (tysm, eqs))) = Fun (c, (tysm, (map o apfst)
+ (fn (ts, t) => (map f ts, f t)) eqs))
+ | map_terms_stmt f (stmt as Datatype _) = stmt
+ | map_terms_stmt f (stmt as Datatypecons _) = stmt
+ | map_terms_stmt f (stmt as Class _) = stmt
+ | map_terms_stmt f (stmt as Classrel _) = stmt
+ | map_terms_stmt f (stmt as Classparam _) = stmt
+ | map_terms_stmt f (Classinst (arity, (superarities, classparms))) =
+ Classinst (arity, (superarities, (map o apfst o apsnd) (fn const =>
+ case f (IConst const) of IConst const' => const') classparms));
+
+fun is_cons program name = case Graph.get_node program name
+ of Datatypecons _ => true
+ | _ => false;
+
+fun contr_classparam_typs program name = case Graph.get_node program name
+ of Classparam (_, class) => let
+ val Class (_, (_, (_, params))) = Graph.get_node program class;
+ val SOME ty = AList.lookup (op =) params name;
+ val (tys, res_ty) = unfold_fun ty;
+ fun no_tyvar (_ `%% tys) = forall no_tyvar tys
+ | no_tyvar (ITyVar _) = false;
+ in if no_tyvar res_ty
+ then map (fn ty => if no_tyvar ty then NONE else SOME ty) tys
+ else []
+ end
+ | _ => [];
+
+
+(** translation kernel **)
+
+(* generic mechanisms *)
+
+fun ensure_stmt lookup declare generate thing (dep, (naming, program)) =
+ let
+ fun add_dep name = case dep of NONE => I
+ | SOME dep => Graph.add_edge (dep, name);
+ val (name, naming') = case lookup naming thing
+ of SOME name => (name, naming)
+ | NONE => declare thing naming;
+ in case try (Graph.get_node program) name
+ of SOME stmt => program
+ |> add_dep name
+ |> pair naming'
+ |> pair dep
+ |> pair name
+ | NONE => program
+ |> Graph.default_node (name, NoStmt)
+ |> add_dep name
+ |> pair naming'
+ |> curry generate (SOME name)
+ ||> snd
+ |-> (fn stmt => (apsnd o Graph.map_node name) (K stmt))
+ |> pair dep
+ |> pair name
+ end;
+
+fun not_wellsorted thy thm ty sort e =
+ let
+ val err_class = Sorts.class_error (Syntax.pp_global thy) e;
+ val err_thm = case thm
+ of SOME thm => "\n(in code equation " ^ Display.string_of_thm thm ^ ")" | NONE => "";
+ val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort "
+ ^ Syntax.string_of_sort_global thy sort;
+ in error ("Wellsortedness error" ^ err_thm ^ ":\n" ^ err_typ ^ "\n" ^ err_class) end;
+
+
+(* translation *)
+
+fun ensure_tyco thy algbr funcgr tyco =
+ let
+ val stmt_datatype =
+ let
+ val (vs, cos) = Code.get_datatype thy tyco;
+ in
+ fold_map (translate_tyvar_sort thy algbr funcgr) vs
+ ##>> fold_map (fn (c, tys) =>
+ ensure_const thy algbr funcgr c
+ ##>> fold_map (translate_typ thy algbr funcgr) tys) cos
+ #>> (fn info => Datatype (tyco, info))
+ end;
+ in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end
+and ensure_const thy algbr funcgr c =
+ let
+ fun stmt_datatypecons tyco =
+ ensure_tyco thy algbr funcgr tyco
+ #>> (fn tyco => Datatypecons (c, tyco));
+ fun stmt_classparam class =
+ ensure_class thy algbr funcgr class
+ #>> (fn class => Classparam (c, class));
+ fun stmt_fun ((vs, ty), raw_thms) =
+ let
+ val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
+ then raw_thms
+ else (map o apfst) (Code.expand_eta thy 1) raw_thms;
+ in
+ fold_map (translate_tyvar_sort thy algbr funcgr) vs
+ ##>> translate_typ thy algbr funcgr ty
+ ##>> fold_map (translate_eq thy algbr funcgr) thms
+ #>> (fn info => Fun (c, info))
+ end;
+ val stmt_const = case Code.get_datatype_of_constr thy c
+ of SOME tyco => stmt_datatypecons tyco
+ | NONE => (case AxClass.class_of_param thy c
+ of SOME class => stmt_classparam class
+ | NONE => stmt_fun (Code_Preproc.typ funcgr c, Code_Preproc.eqns funcgr c))
+ in ensure_stmt lookup_const (declare_const thy) stmt_const c end
+and ensure_class thy (algbr as (_, algebra)) funcgr class =
+ let
+ val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
+ val cs = #params (AxClass.get_info thy class);
+ val stmt_class =
+ fold_map (fn superclass => ensure_class thy algbr funcgr superclass
+ ##>> ensure_classrel thy algbr funcgr (class, superclass)) superclasses
+ ##>> fold_map (fn (c, ty) => ensure_const thy algbr funcgr c
+ ##>> translate_typ thy algbr funcgr ty) cs
+ #>> (fn info => Class (class, (unprefix "'" Name.aT, info)))
+ in ensure_stmt lookup_class (declare_class thy) stmt_class class end
+and ensure_classrel thy algbr funcgr (subclass, superclass) =
+ let
+ val stmt_classrel =
+ ensure_class thy algbr funcgr subclass
+ ##>> ensure_class thy algbr funcgr superclass
+ #>> Classrel;
+ in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (subclass, superclass) end
+and ensure_inst thy (algbr as (_, algebra)) funcgr (class, tyco) =
+ let
+ val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
+ val classparams = these (try (#params o AxClass.get_info thy) class);
+ val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]);
+ val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class];
+ val vs' = map2 (fn (v, sort1) => fn sort2 => (v,
+ Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts';
+ val arity_typ = Type (tyco, map TFree vs);
+ val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs');
+ fun translate_superarity superclass =
+ ensure_class thy algbr funcgr superclass
+ ##>> ensure_classrel thy algbr funcgr (class, superclass)
+ ##>> translate_dicts thy algbr funcgr NONE (arity_typ, [superclass])
+ #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
+ (superclass, (classrel, (inst, dss))));
+ fun translate_classparam_inst (c, ty) =
+ let
+ val c_inst = Const (c, map_type_tfree (K arity_typ') ty);
+ val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy c_inst);
+ val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd
+ o Logic.dest_equals o Thm.prop_of) thm;
+ in
+ ensure_const thy algbr funcgr c
+ ##>> translate_const thy algbr funcgr (SOME thm) c_ty
+ #>> (fn (c, IConst c_inst) => ((c, c_inst), (thm, true)))
+ end;
+ val stmt_inst =
+ ensure_class thy algbr funcgr class
+ ##>> ensure_tyco thy algbr funcgr tyco
+ ##>> fold_map (translate_tyvar_sort thy algbr funcgr) vs
+ ##>> fold_map translate_superarity superclasses
+ ##>> fold_map translate_classparam_inst classparams
+ #>> (fn ((((class, tyco), arity), superarities), classparams) =>
+ Classinst ((class, (tyco, arity)), (superarities, classparams)));
+ in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end
+and translate_typ thy algbr funcgr (TFree (v, _)) =
+ pair (ITyVar (unprefix "'" v))
+ | translate_typ thy algbr funcgr (Type (tyco, tys)) =
+ ensure_tyco thy algbr funcgr tyco
+ ##>> fold_map (translate_typ thy algbr funcgr) tys
+ #>> (fn (tyco, tys) => tyco `%% tys)
+and translate_term thy algbr funcgr thm (Const (c, ty)) =
+ translate_app thy algbr funcgr thm ((c, ty), [])
+ | translate_term thy algbr funcgr thm (Free (v, _)) =
+ pair (IVar v)
+ | translate_term thy algbr funcgr thm (Abs (abs as (_, ty, _))) =
+ let
+ val (v, t) = Syntax.variant_abs abs;
+ in
+ translate_typ thy algbr funcgr ty
+ ##>> translate_term thy algbr funcgr thm t
+ #>> (fn (ty, t) => (v, ty) `|=> t)
+ end
+ | translate_term thy algbr funcgr thm (t as _ $ _) =
+ case strip_comb t
+ of (Const (c, ty), ts) =>
+ translate_app thy algbr funcgr thm ((c, ty), ts)
+ | (t', ts) =>
+ translate_term thy algbr funcgr thm t'
+ ##>> fold_map (translate_term thy algbr funcgr thm) ts
+ #>> (fn (t, ts) => t `$$ ts)
+and translate_eq thy algbr funcgr (thm, proper) =
+ let
+ val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals
+ o Logic.unvarify o prop_of) thm;
+ in
+ fold_map (translate_term thy algbr funcgr (SOME thm)) args
+ ##>> translate_term thy algbr funcgr (SOME thm) rhs
+ #>> rpair (thm, proper)
+ end
+and translate_const thy algbr funcgr thm (c, ty) =
+ let
+ val tys = Sign.const_typargs thy (c, ty);
+ val sorts = (map snd o fst o Code_Preproc.typ funcgr) c;
+ val tys_args = (fst o Term.strip_type) ty;
+ in
+ ensure_const thy algbr funcgr c
+ ##>> fold_map (translate_typ thy algbr funcgr) tys
+ ##>> fold_map (translate_dicts thy algbr funcgr thm) (tys ~~ sorts)
+ ##>> fold_map (translate_typ thy algbr funcgr) tys_args
+ #>> (fn (((c, tys), iss), tys_args) => IConst (c, ((tys, iss), tys_args)))
+ end
+and translate_app_const thy algbr funcgr thm (c_ty, ts) =
+ translate_const thy algbr funcgr thm c_ty
+ ##>> fold_map (translate_term thy algbr funcgr thm) ts
+ #>> (fn (t, ts) => t `$$ ts)
+and translate_case thy algbr funcgr thm (num_args, (t_pos, case_pats)) (c_ty, ts) =
+ let
+ val (tys, _) = (chop num_args o fst o strip_type o snd) c_ty;
+ val t = nth ts t_pos;
+ val ty = nth tys t_pos;
+ val ts_clause = nth_drop t_pos ts;
+ fun mk_clause (co, num_co_args) t =
+ let
+ val (vs, body) = Term.strip_abs_eta num_co_args t;
+ val not_undefined = case body
+ of (Const (c, _)) => not (Code.is_undefined thy c)
+ | _ => true;
+ val pat = list_comb (Const (co, map snd vs ---> ty), map Free vs);
+ in (not_undefined, (pat, body)) end;
+ val clauses = if null case_pats then let val ([v_ty], body) =
+ Term.strip_abs_eta 1 (the_single ts_clause)
+ in [(true, (Free v_ty, body))] end
+ else map (uncurry mk_clause)
+ (AList.make (Code.no_args thy) case_pats ~~ ts_clause);
+ fun retermify ty (_, (IVar x, body)) =
+ (x, ty) `|=> body
+ | retermify _ (_, (pat, body)) =
+ let
+ val (IConst (_, (_, tys)), ts) = unfold_app pat;
+ val vs = map2 (fn IVar x => fn ty => (x, ty)) ts tys;
+ in vs `|==> body end;
+ fun mk_icase const t ty clauses =
+ let
+ val (ts1, ts2) = chop t_pos (map (retermify ty) clauses);
+ in
+ ICase (((t, ty), map_filter (fn (b, d) => if b then SOME d else NONE) clauses),
+ const `$$ (ts1 @ t :: ts2))
+ end;
+ in
+ translate_const thy algbr funcgr thm c_ty
+ ##>> translate_term thy algbr funcgr thm t
+ ##>> translate_typ thy algbr funcgr ty
+ ##>> fold_map (fn (b, (pat, body)) => translate_term thy algbr funcgr thm pat
+ ##>> translate_term thy algbr funcgr thm body
+ #>> pair b) clauses
+ #>> (fn (((const, t), ty), ds) => mk_icase const t ty ds)
+ end
+and translate_app_case thy algbr funcgr thm (case_scheme as (num_args, _)) ((c, ty), ts) =
+ if length ts < num_args then
+ let
+ val k = length ts;
+ val tys = (curry Library.take (num_args - k) o curry Library.drop k o fst o strip_type) ty;
+ val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context;
+ val vs = Name.names ctxt "a" tys;
+ in
+ fold_map (translate_typ thy algbr funcgr) tys
+ ##>> translate_case thy algbr funcgr thm case_scheme ((c, ty), ts @ map Free vs)
+ #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|==> t)
+ end
+ else if length ts > num_args then
+ translate_case thy algbr funcgr thm case_scheme ((c, ty), Library.take (num_args, ts))
+ ##>> fold_map (translate_term thy algbr funcgr thm) (Library.drop (num_args, ts))
+ #>> (fn (t, ts) => t `$$ ts)
+ else
+ translate_case thy algbr funcgr thm case_scheme ((c, ty), ts)
+and translate_app thy algbr funcgr thm (c_ty_ts as ((c, _), _)) =
+ case Code.get_case_scheme thy c
+ of SOME case_scheme => translate_app_case thy algbr funcgr thm case_scheme c_ty_ts
+ | NONE => translate_app_const thy algbr funcgr thm c_ty_ts
+and translate_tyvar_sort thy (algbr as (proj_sort, _)) funcgr (v, sort) =
+ fold_map (ensure_class thy algbr funcgr) (proj_sort sort)
+ #>> (fn sort => (unprefix "'" v, sort))
+and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) =
+ let
+ val pp = Syntax.pp_global thy;
+ datatype typarg =
+ Global of (class * string) * typarg list list
+ | Local of (class * class) list * (string * (int * sort));
+ fun class_relation (Global ((_, tyco), yss), _) class =
+ Global ((class, tyco), yss)
+ | class_relation (Local (classrels, v), subclass) superclass =
+ Local ((subclass, superclass) :: classrels, v);
+ fun type_constructor tyco yss class =
+ Global ((class, tyco), (map o map) fst yss);
+ fun type_variable (TFree (v, sort)) =
+ let
+ val sort' = proj_sort sort;
+ in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
+ val typargs = Sorts.of_sort_derivation pp algebra
+ {class_relation = class_relation, type_constructor = type_constructor,
+ type_variable = type_variable} (ty, proj_sort sort)
+ handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e;
+ fun mk_dict (Global (inst, yss)) =
+ ensure_inst thy algbr funcgr inst
+ ##>> (fold_map o fold_map) mk_dict yss
+ #>> (fn (inst, dss) => DictConst (inst, dss))
+ | mk_dict (Local (classrels, (v, (k, sort)))) =
+ fold_map (ensure_classrel thy algbr funcgr) classrels
+ #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
+ in fold_map mk_dict typargs end;
+
+
+(* store *)
+
+structure Program = CodeDataFun
+(
+ type T = naming * program;
+ val empty = (empty_naming, Graph.empty);
+ fun purge thy cs (naming, program) =
+ let
+ val names_delete = cs
+ |> map_filter (lookup_const naming)
+ |> filter (can (Graph.get_node program))
+ |> Graph.all_preds program;
+ val program' = Graph.del_nodes names_delete program;
+ in (naming, program') end;
+);
+
+val cached_program = Program.get;
+
+fun invoke_generation thy (algebra, funcgr) f name =
+ Program.change_yield thy (fn naming_program => (NONE, naming_program)
+ |> f thy algebra funcgr name
+ |-> (fn name => fn (_, naming_program) => (name, naming_program)));
+
+
+(* program generation *)
+
+fun consts_program thy cs =
+ let
+ fun project_consts cs (naming, program) =
+ let
+ val cs_all = Graph.all_succs program cs;
+ in (cs, (naming, Graph.subgraph (member (op =) cs_all) program)) end;
+ fun generate_consts thy algebra funcgr =
+ fold_map (ensure_const thy algebra funcgr);
+ in
+ invoke_generation thy (Code_Preproc.obtain thy cs []) generate_consts cs
+ |-> project_consts
+ end;
+
+
+(* value evaluation *)
+
+fun ensure_value thy algbr funcgr t =
+ let
+ val ty = fastype_of t;
+ val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =)
+ o dest_TFree))) t [];
+ val stmt_value =
+ fold_map (translate_tyvar_sort thy algbr funcgr) vs
+ ##>> translate_typ thy algbr funcgr ty
+ ##>> translate_term thy algbr funcgr NONE t
+ #>> (fn ((vs, ty), t) => Fun
+ (Term.dummy_patternN, ((vs, ty), [(([], t), (Drule.dummy_thm, true))])));
+ fun term_value (dep, (naming, program1)) =
+ let
+ val Fun (_, (vs_ty, [(([], t), _)])) =
+ Graph.get_node program1 Term.dummy_patternN;
+ val deps = Graph.imm_succs program1 Term.dummy_patternN;
+ val program2 = Graph.del_nodes [Term.dummy_patternN] program1;
+ val deps_all = Graph.all_succs program2 deps;
+ val program3 = Graph.subgraph (member (op =) deps_all) program2;
+ in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end;
+ in
+ ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN
+ #> snd
+ #> term_value
+ end;
+
+fun base_evaluator thy evaluator algebra funcgr vs t =
+ let
+ val (((naming, program), (((vs', ty'), t'), deps)), _) =
+ invoke_generation thy (algebra, funcgr) ensure_value t;
+ val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';
+ in evaluator naming program ((vs'', (vs', ty')), t') deps end;
+
+fun eval_conv thy prep_sort = Code_Preproc.eval_conv thy prep_sort o base_evaluator thy;
+fun eval thy prep_sort postproc = Code_Preproc.eval thy prep_sort postproc o base_evaluator thy;
+
+
+(** diagnostic commands **)
+
+fun read_const_exprs thy =
+ let
+ fun consts_of some_thyname =
+ let
+ val thy' = case some_thyname
+ of SOME thyname => ThyInfo.the_theory thyname thy
+ | NONE => thy;
+ val cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I)
+ ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) thy') [];
+ fun belongs_here c =
+ not (exists (fn thy'' => Sign.declared_const thy'' c) (Theory.parents_of thy'))
+ in case some_thyname
+ of NONE => cs
+ | SOME thyname => filter belongs_here cs
+ end;
+ fun read_const_expr "*" = ([], consts_of NONE)
+ | read_const_expr s = if String.isSuffix ".*" s
+ then ([], consts_of (SOME (unsuffix ".*" s)))
+ else ([Code.read_const thy s], []);
+ in pairself flat o split_list o map read_const_expr end;
+
+fun code_depgr thy consts =
+ let
+ val (_, eqngr) = Code_Preproc.obtain thy consts [];
+ val select = Graph.all_succs eqngr consts;
+ in
+ eqngr
+ |> not (null consts) ? Graph.subgraph (member (op =) select)
+ |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy))
+ end;
+
+fun code_thms thy = Pretty.writeln o Code_Preproc.pretty thy o code_depgr thy;
+
+fun code_deps thy consts =
+ let
+ val eqngr = code_depgr thy consts;
+ val constss = Graph.strong_conn eqngr;
+ val mapping = Symtab.empty |> fold (fn consts => fold (fn const =>
+ Symtab.update (const, consts)) consts) constss;
+ fun succs consts = consts
+ |> maps (Graph.imm_succs eqngr)
+ |> subtract (op =) consts
+ |> map (the o Symtab.lookup mapping)
+ |> distinct (op =);
+ val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
+ fun namify consts = map (Code.string_of_const thy) consts
+ |> commas;
+ val prgr = map (fn (consts, constss) =>
+ { name = namify consts, ID = namify consts, dir = "", unfold = true,
+ path = "", parents = map namify constss }) conn;
+ in Present.display_graph prgr end;
+
+local
+
+structure P = OuterParse
+and K = OuterKeyword
+
+fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy;
+fun code_deps_cmd thy = code_deps thy o op @ o read_const_exprs thy;
+
+in
+
+val _ =
+ OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag
+ (Scan.repeat P.term_group
+ >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
+ o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
+
+val _ =
+ OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag
+ (Scan.repeat P.term_group
+ >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
+ o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
+
+end;
+
+end; (*struct*)
+
+
+structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol;
--- a/src/Tools/Code_Generator.thy Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Tools/Code_Generator.thy Tue Jun 23 15:48:55 2009 +0100
@@ -9,12 +9,12 @@
uses
"~~/src/Tools/value.ML"
"~~/src/Tools/quickcheck.ML"
- "~~/src/Tools/code/code_preproc.ML"
- "~~/src/Tools/code/code_thingol.ML"
- "~~/src/Tools/code/code_printer.ML"
- "~~/src/Tools/code/code_target.ML"
- "~~/src/Tools/code/code_ml.ML"
- "~~/src/Tools/code/code_haskell.ML"
+ "~~/src/Tools/Code/code_preproc.ML"
+ "~~/src/Tools/Code/code_thingol.ML"
+ "~~/src/Tools/Code/code_printer.ML"
+ "~~/src/Tools/Code/code_target.ML"
+ "~~/src/Tools/Code/code_ml.ML"
+ "~~/src/Tools/Code/code_haskell.ML"
"~~/src/Tools/nbe.ML"
begin
--- a/src/Tools/code/code_haskell.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,567 +0,0 @@
-(* Title: Tools/code/code_haskell.ML
- Author: Florian Haftmann, TU Muenchen
-
-Serializer for Haskell.
-*)
-
-signature CODE_HASKELL =
-sig
- val setup: theory -> theory
-end;
-
-structure Code_Haskell : CODE_HASKELL =
-struct
-
-val target = "Haskell";
-
-open Basic_Code_Thingol;
-open Code_Printer;
-
-infixr 5 @@;
-infixr 5 @|;
-
-
-(** Haskell serializer **)
-
-fun pr_haskell_bind pr_term =
- let
- fun pr_bind ((NONE, NONE), _) = str "_"
- | pr_bind ((SOME v, NONE), _) = str v
- | pr_bind ((NONE, SOME p), _) = p
- | pr_bind ((SOME v, SOME p), _) = brackets [str v, str "@", p];
- in gen_pr_bind pr_bind pr_term end;
-
-fun pr_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const
- init_syms deresolve is_cons contr_classparam_typs deriving_show =
- let
- val deresolve_base = Long_Name.base_name o deresolve;
- fun class_name class = case syntax_class class
- of NONE => deresolve class
- | SOME class => class;
- fun pr_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
- of [] => []
- | classbinds => Pretty.enum "," "(" ")" (
- map (fn (v, class) =>
- str (class_name class ^ " " ^ Code_Printer.lookup_var tyvars v)) classbinds)
- @@ str " => ";
- fun pr_typforall tyvars vs = case map fst vs
- of [] => []
- | vnames => str "forall " :: Pretty.breaks
- (map (str o Code_Printer.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
- fun pr_tycoexpr tyvars fxy (tyco, tys) =
- brackify fxy (str tyco :: map (pr_typ tyvars BR) tys)
- and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco
- of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys)
- | SOME (i, pr) => pr (pr_typ tyvars) fxy tys)
- | pr_typ tyvars fxy (ITyVar v) = (str o Code_Printer.lookup_var tyvars) v;
- fun pr_typdecl tyvars (vs, tycoexpr) =
- Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr);
- fun pr_typscheme tyvars (vs, ty) =
- Pretty.block (pr_typforall tyvars vs @ pr_typcontext tyvars vs @| pr_typ tyvars NOBR ty);
- fun pr_term tyvars thm vars fxy (IConst c) =
- pr_app tyvars thm vars fxy (c, [])
- | pr_term tyvars thm vars fxy (t as (t1 `$ t2)) =
- (case Code_Thingol.unfold_const_app t
- of SOME app => pr_app tyvars thm vars fxy app
- | _ =>
- brackify fxy [
- pr_term tyvars thm vars NOBR t1,
- pr_term tyvars thm vars BR t2
- ])
- | pr_term tyvars thm vars fxy (IVar v) =
- (str o Code_Printer.lookup_var vars) v
- | pr_term tyvars thm vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = Code_Thingol.unfold_abs t;
- fun pr ((v, pat), ty) = pr_bind tyvars thm BR ((SOME v, pat), ty);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (str "\\" :: ps @ str "->" @@ pr_term tyvars thm vars' NOBR t') end
- | pr_term tyvars thm vars fxy (ICase (cases as (_, t0))) =
- (case Code_Thingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
- then pr_case tyvars thm vars fxy cases
- else pr_app tyvars thm vars fxy c_ts
- | NONE => pr_case tyvars thm vars fxy cases)
- and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c
- of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts
- | fingerprint => let
- val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint;
- val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) =>
- (not o Code_Thingol.locally_monomorphic) t) ts_fingerprint;
- fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t
- | pr_term_anno (t, SOME _) ty =
- brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty];
- in
- if needs_annotation then
- (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys)
- else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts
- end
- and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const
- and pr_bind tyvars = pr_haskell_bind (pr_term tyvars)
- and pr_case tyvars thm vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, body) = Code_Thingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind tyvars thm BR ((NONE, SOME pat), ty)
- |>> (fn p => semicolon [p, str "=", pr_term tyvars thm vars NOBR t])
- val (ps, vars') = fold_map pr binds vars;
- in brackify_block fxy (str "let {")
- ps
- (concat [str "}", str "in", pr_term tyvars thm vars' NOBR body])
- end
- | pr_case tyvars thm vars fxy (((t, ty), clauses as _ :: _), _) =
- let
- fun pr (pat, body) =
- let
- val (p, vars') = pr_bind tyvars thm NOBR ((NONE, SOME pat), ty) vars;
- in semicolon [p, str "->", pr_term tyvars thm vars' NOBR body] end;
- in brackify_block fxy
- (concat [str "case", pr_term tyvars thm vars NOBR t, str "of", str "{"])
- (map pr clauses)
- (str "}")
- end
- | pr_case tyvars thm vars fxy ((_, []), _) =
- (brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""];
- fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) =
- let
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- val n = (length o fst o Code_Thingol.unfold_fun) ty;
- in
- Pretty.chunks [
- Pretty.block [
- (str o suffix " ::" o deresolve_base) name,
- Pretty.brk 1,
- pr_typscheme tyvars (vs, ty),
- str ";"
- ],
- concat (
- (str o deresolve_base) name
- :: map str (replicate n "_")
- @ str "="
- :: str "error"
- @@ (str o (fn s => s ^ ";") o ML_Syntax.print_string
- o Long_Name.base_name o Long_Name.qualifier) name
- )
- ]
- end
- | pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) =
- let
- val eqs = filter (snd o snd) raw_eqs;
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- fun pr_eq ((ts, t), (thm, _)) =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = init_syms
- |> Code_Printer.intro_vars consts
- |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- semicolon (
- (str o deresolve_base) name
- :: map (pr_term tyvars thm vars BR) ts
- @ str "="
- @@ pr_term tyvars thm vars NOBR t
- )
- end;
- in
- Pretty.chunks (
- Pretty.block [
- (str o suffix " ::" o deresolve_base) name,
- Pretty.brk 1,
- pr_typscheme tyvars (vs, ty),
- str ";"
- ]
- :: map pr_eq eqs
- )
- end
- | pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) =
- let
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- in
- semicolon [
- str "data",
- pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
- ]
- end
- | pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) =
- let
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- in
- semicolon (
- str "newtype"
- :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
- :: str "="
- :: (str o deresolve_base) co
- :: pr_typ tyvars BR ty
- :: (if deriving_show name then [str "deriving (Read, Show)"] else [])
- )
- end
- | pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) =
- let
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- fun pr_co (co, tys) =
- concat (
- (str o deresolve_base) co
- :: map (pr_typ tyvars BR) tys
- )
- in
- semicolon (
- str "data"
- :: pr_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
- :: str "="
- :: pr_co co
- :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos
- @ (if deriving_show name then [str "deriving (Read, Show)"] else [])
- )
- end
- | pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) =
- let
- val tyvars = Code_Printer.intro_vars [v] init_syms;
- fun pr_classparam (classparam, ty) =
- semicolon [
- (str o deresolve_base) classparam,
- str "::",
- pr_typ tyvars NOBR ty
- ]
- in
- Pretty.block_enclose (
- Pretty.block [
- str "class ",
- Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]),
- str (deresolve_base name ^ " " ^ Code_Printer.lookup_var tyvars v),
- str " where {"
- ],
- str "};"
- ) (map pr_classparam classparams)
- end
- | pr_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) =
- let
- val split_abs_pure = (fn (v, _) `|-> t => SOME (v, t) | _ => NONE);
- val unfold_abs_pure = Code_Thingol.unfoldr split_abs_pure;
- val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
- fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam
- of NONE => semicolon [
- (str o deresolve_base) classparam,
- str "=",
- pr_app tyvars thm init_syms NOBR (c_inst, [])
- ]
- | SOME (k, pr) =>
- let
- val (c_inst_name, (_, tys)) = c_inst;
- val const = if (is_some o syntax_const) c_inst_name
- then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name;
- val proto_rhs = Code_Thingol.eta_expand k (c_inst, []);
- val (vs, rhs) = unfold_abs_pure proto_rhs;
- val vars = init_syms
- |> Code_Printer.intro_vars (the_list const)
- |> Code_Printer.intro_vars vs;
- val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs;
- (*dictionaries are not relevant at this late stage*)
- in
- semicolon [
- pr_term tyvars thm vars NOBR lhs,
- str "=",
- pr_term tyvars thm vars NOBR rhs
- ]
- end;
- in
- Pretty.block_enclose (
- Pretty.block [
- str "instance ",
- Pretty.block (pr_typcontext tyvars vs),
- str (class_name class ^ " "),
- pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
- str " where {"
- ],
- str "};"
- ) (map pr_instdef classparam_insts)
- end;
- in pr_stmt end;
-
-fun haskell_program_of_program labelled_name module_name module_prefix reserved_names raw_module_alias program =
- let
- val module_alias = if is_some module_name then K module_name else raw_module_alias;
- val reserved_names = Name.make_context reserved_names;
- val mk_name_module = Code_Printer.mk_name_module reserved_names module_prefix module_alias program;
- fun add_stmt (name, (stmt, deps)) =
- let
- val (module_name, base) = Code_Printer.dest_name name;
- val module_name' = mk_name_module module_name;
- val mk_name_stmt = yield_singleton Name.variants;
- fun add_fun upper (nsp_fun, nsp_typ) =
- let
- val (base', nsp_fun') =
- mk_name_stmt (if upper then Code_Printer.first_upper base else base) nsp_fun
- in (base', (nsp_fun', nsp_typ)) end;
- fun add_typ (nsp_fun, nsp_typ) =
- let
- val (base', nsp_typ') = mk_name_stmt (Code_Printer.first_upper base) nsp_typ
- in (base', (nsp_fun, nsp_typ')) end;
- val add_name = case stmt
- of Code_Thingol.Fun _ => add_fun false
- | Code_Thingol.Datatype _ => add_typ
- | Code_Thingol.Datatypecons _ => add_fun true
- | Code_Thingol.Class _ => add_typ
- | Code_Thingol.Classrel _ => pair base
- | Code_Thingol.Classparam _ => add_fun false
- | Code_Thingol.Classinst _ => pair base;
- fun add_stmt' base' = case stmt
- of Code_Thingol.Datatypecons _ =>
- cons (name, (Long_Name.append module_name' base', NONE))
- | Code_Thingol.Classrel _ => I
- | Code_Thingol.Classparam _ =>
- cons (name, (Long_Name.append module_name' base', NONE))
- | _ => cons (name, (Long_Name.append module_name' base', SOME stmt));
- in
- Symtab.map_default (module_name', ([], ([], (reserved_names, reserved_names))))
- (apfst (fold (insert (op = : string * string -> bool)) deps))
- #> `(fn program => add_name ((snd o snd o the o Symtab.lookup program) module_name'))
- #-> (fn (base', names) =>
- (Symtab.map_entry module_name' o apsnd) (fn (stmts, _) =>
- (add_stmt' base' stmts, names)))
- end;
- val hs_program = fold add_stmt (AList.make (fn name =>
- (Graph.get_node program name, Graph.imm_succs program name))
- (Graph.strong_conn program |> flat)) Symtab.empty;
- fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the
- o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Printer.dest_name) name))) name
- handle Option => error ("Unknown statement name: " ^ labelled_name name);
- in (deresolver, hs_program) end;
-
-fun serialize_haskell module_prefix raw_module_name string_classes labelled_name
- raw_reserved_names includes raw_module_alias
- syntax_class syntax_tyco syntax_const program cs destination =
- let
- val stmt_names = Code_Target.stmt_names_of_destination destination;
- val module_name = if null stmt_names then raw_module_name else SOME "Code";
- val reserved_names = fold (insert (op =) o fst) includes raw_reserved_names;
- val (deresolver, hs_program) = haskell_program_of_program labelled_name
- module_name module_prefix reserved_names raw_module_alias program;
- val is_cons = Code_Thingol.is_cons program;
- val contr_classparam_typs = Code_Thingol.contr_classparam_typs program;
- fun deriving_show tyco =
- let
- fun deriv _ "fun" = false
- | deriv tycos tyco = member (op =) tycos tyco orelse
- case try (Graph.get_node program) tyco
- of SOME (Code_Thingol.Datatype (_, (_, cs))) => forall (deriv' (tyco :: tycos))
- (maps snd cs)
- | NONE => true
- and deriv' tycos (tyco `%% tys) = deriv tycos tyco
- andalso forall (deriv' tycos) tys
- | deriv' _ (ITyVar _) = true
- in deriv [] tyco end;
- val reserved_names = Code_Printer.make_vars reserved_names;
- fun pr_stmt qualified = pr_haskell_stmt labelled_name
- syntax_class syntax_tyco syntax_const reserved_names
- (if qualified then deresolver else Long_Name.base_name o deresolver)
- is_cons contr_classparam_typs
- (if string_classes then deriving_show else K false);
- fun pr_module name content =
- (name, Pretty.chunks [
- str ("module " ^ name ^ " where {"),
- str "",
- content,
- str "",
- str "}"
- ]);
- fun serialize_module1 (module_name', (deps, (stmts, _))) =
- let
- val stmt_names = map fst stmts;
- val deps' = subtract (op =) stmt_names deps
- |> distinct (op =)
- |> map_filter (try deresolver);
- val qualified = is_none module_name andalso
- map deresolver stmt_names @ deps'
- |> map Long_Name.base_name
- |> has_duplicates (op =);
- val imports = deps'
- |> map Long_Name.qualifier
- |> distinct (op =);
- fun pr_import_include (name, _) = str ("import qualified " ^ name ^ ";");
- val pr_import_module = str o (if qualified
- then prefix "import qualified "
- else prefix "import ") o suffix ";";
- val content = Pretty.chunks (
- map pr_import_include includes
- @ map pr_import_module imports
- @ str ""
- :: separate (str "") (map_filter
- (fn (name, (_, SOME stmt)) => SOME (pr_stmt qualified (name, stmt))
- | (_, (_, NONE)) => NONE) stmts)
- )
- in pr_module module_name' content end;
- fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks (
- separate (str "") (map_filter
- (fn (name, (_, SOME stmt)) => if null stmt_names
- orelse member (op =) stmt_names name
- then SOME (pr_stmt false (name, stmt))
- else NONE
- | (_, (_, NONE)) => NONE) stmts));
- val serialize_module =
- if null stmt_names then serialize_module1 else pair "" o serialize_module2;
- fun check_destination destination =
- (File.check destination; destination);
- fun write_module destination (modlname, content) =
- let
- val filename = case modlname
- of "" => Path.explode "Main.hs"
- | _ => (Path.ext "hs" o Path.explode o implode o separate "/"
- o Long_Name.explode) modlname;
- val pathname = Path.append destination filename;
- val _ = File.mkdir (Path.dir pathname);
- in File.write pathname
- ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n"
- ^ Code_Target.code_of_pretty content)
- end
- in
- Code_Target.mk_serialization target NONE
- (fn NONE => K () o map (Code_Target.code_writeln o snd) | SOME file => K () o map
- (write_module (check_destination file)))
- (rpair [] o cat_lines o map (Code_Target.code_of_pretty o snd))
- (map (uncurry pr_module) includes
- @ map serialize_module (Symtab.dest hs_program))
- destination
- end;
-
-val literals = let
- fun char_haskell c =
- let
- val s = ML_Syntax.print_char c;
- in if s = "'" then "\\'" else s end;
-in Literals {
- literal_char = enclose "'" "'" o char_haskell,
- literal_string = quote o translate_string char_haskell,
- literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k
- else enclose "(" ")" (signed_string_of_int k),
- literal_list = Pretty.enum "," "[" "]",
- infix_cons = (5, ":")
-} end;
-
-
-(** optional monad syntax **)
-
-fun pretty_haskell_monad c_bind =
- let
- fun dest_bind t1 t2 = case Code_Thingol.split_abs t2
- of SOME (((v, pat), ty), t') =>
- SOME ((SOME (((SOME v, pat), ty), true), t1), t')
- | NONE => NONE;
- fun dest_monad c_bind_name (IConst (c, _) `$ t1 `$ t2) =
- if c = c_bind_name then dest_bind t1 t2
- else NONE
- | dest_monad _ t = case Code_Thingol.split_let t
- of SOME (((pat, ty), tbind), t') =>
- SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t')
- | NONE => NONE;
- fun implode_monad c_bind_name = Code_Thingol.unfoldr (dest_monad c_bind_name);
- fun pr_monad pr_bind pr (NONE, t) vars =
- (semicolon [pr vars NOBR t], vars)
- | pr_monad pr_bind pr (SOME (bind, true), t) vars = vars
- |> pr_bind NOBR bind
- |>> (fn p => semicolon [p, str "<-", pr vars NOBR t])
- | pr_monad pr_bind pr (SOME (bind, false), t) vars = vars
- |> pr_bind NOBR bind
- |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]);
- fun pretty _ [c_bind'] pr thm vars fxy [(t1, _), (t2, _)] = case dest_bind t1 t2
- of SOME (bind, t') => let
- val (binds, t'') = implode_monad c_bind' t'
- val (ps, vars') = fold_map (pr_monad (pr_haskell_bind (K pr) thm) pr) (bind :: binds) vars;
- in (brackify fxy o single o Pretty.enclose "do {" "}" o Pretty.breaks) (ps @| pr vars' NOBR t'') end
- | NONE => brackify_infix (1, L) fxy
- [pr vars (INFX (1, L)) t1, str ">>=", pr vars (INFX (1, X)) t2]
- in (2, ([c_bind], pretty)) end;
-
-fun add_monad target' raw_c_bind thy =
- let
- val c_bind = Code.read_const thy raw_c_bind;
- in if target = target' then
- thy
- |> Code_Target.add_syntax_const target c_bind
- (SOME (pretty_haskell_monad c_bind))
- else error "Only Haskell target allows for monad syntax" end;
-
-
-(** Isar setup **)
-
-fun isar_seri_haskell module =
- Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
- -- Scan.optional (Args.$$$ "string_classes" >> K true) false
- >> (fn (module_prefix, string_classes) =>
- serialize_haskell module_prefix module string_classes));
-
-val _ =
- OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl (
- OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) =>
- Toplevel.theory (add_monad target raw_bind))
- );
-
-val setup =
- Code_Target.add_target (target, (isar_seri_haskell, literals))
- #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> fold (Code_Target.add_reserved target) [
- "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
- "import", "default", "forall", "let", "in", "class", "qualified", "data",
- "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
- ]
- #> fold (Code_Target.add_reserved target) [
- "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int",
- "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded",
- "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor",
- "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException",
- "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException",
- "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure",
- "ExitSuccess", "False", "GT", "HeapOverflow",
- "IOError", "IOException", "IllegalOperation",
- "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError",
- "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow",
- "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError",
- "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow",
- "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow",
- "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all",
- "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan",
- "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen",
- "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break",
- "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const",
- "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod",
- "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem",
- "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo",
- "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip",
- "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational",
- "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble",
- "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj",
- "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head",
- "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha",
- "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite",
- "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'",
- "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log",
- "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum",
- "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem",
- "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo",
- "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred",
- "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn",
- "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat",
- "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile",
- "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn",
- "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational",
- "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse",
- "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence",
- "sequence_", "show", "showChar", "showException", "showField", "showList",
- "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand",
- "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract",
- "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult",
- "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry",
- "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords",
- "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3"
- ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*);
-
-end; (*struct*)
--- a/src/Tools/code/code_ml.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1122 +0,0 @@
-(* Title: Tools/code/code_ml.ML
- Author: Florian Haftmann, TU Muenchen
-
-Serializer for SML and OCaml.
-*)
-
-signature CODE_ML =
-sig
- val eval: string option -> string * (unit -> 'a) option ref
- -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a
- val target_Eval: string
- val setup: theory -> theory
-end;
-
-structure Code_ML : CODE_ML =
-struct
-
-open Basic_Code_Thingol;
-open Code_Printer;
-
-infixr 5 @@;
-infixr 5 @|;
-
-val target_SML = "SML";
-val target_OCaml = "OCaml";
-val target_Eval = "Eval";
-
-datatype ml_stmt =
- MLExc of string * int
- | MLVal of string * ((typscheme * iterm) * (thm * bool))
- | MLFuns of (string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)) list * string list
- | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list
- | MLClass of string * (vname * ((class * string) list * (string * itype) list))
- | MLClassinst of string * ((class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list));
-
-fun stmt_names_of (MLExc (name, _)) = [name]
- | stmt_names_of (MLVal (name, _)) = [name]
- | stmt_names_of (MLFuns (fs, _)) = map fst fs
- | stmt_names_of (MLDatas ds) = map fst ds
- | stmt_names_of (MLClass (name, _)) = [name]
- | stmt_names_of (MLClassinst (name, _)) = [name];
-
-
-(** SML serailizer **)
-
-fun pr_sml_stmt labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons =
- let
- fun pr_dicts fxy ds =
- let
- fun pr_dictvar (v, (_, 1)) = Code_Printer.first_upper v ^ "_"
- | pr_dictvar (v, (i, _)) = Code_Printer.first_upper v ^ string_of_int (i+1) ^ "_";
- fun pr_proj [] p =
- p
- | pr_proj [p'] p =
- brackets [p', p]
- | pr_proj (ps as _ :: _) p =
- brackets [Pretty.enum " o" "(" ")" ps, p];
- fun pr_dict fxy (DictConst (inst, dss)) =
- brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss)
- | pr_dict fxy (DictVar (classrels, v)) =
- pr_proj (map (str o deresolve) classrels) ((str o pr_dictvar) v)
- in case ds
- of [] => str "()"
- | [d] => pr_dict fxy d
- | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds
- end;
- fun pr_tyvar_dicts vs =
- vs
- |> map (fn (v, sort) => map_index (fn (i, _) =>
- DictVar ([], (v, (i, length sort)))) sort)
- |> map (pr_dicts BR);
- fun pr_tycoexpr fxy (tyco, tys) =
- let
- val tyco' = (str o deresolve) tyco
- in case map (pr_typ BR) tys
- of [] => tyco'
- | [p] => Pretty.block [p, Pretty.brk 1, tyco']
- | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
- end
- and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
- of NONE => pr_tycoexpr fxy (tyco, tys)
- | SOME (i, pr) => pr pr_typ fxy tys)
- | pr_typ fxy (ITyVar v) = str ("'" ^ v);
- fun pr_term is_closure thm vars fxy (IConst c) =
- pr_app is_closure thm vars fxy (c, [])
- | pr_term is_closure thm vars fxy (IVar v) =
- str (Code_Printer.lookup_var vars v)
- | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
- (case Code_Thingol.unfold_const_app t
- of SOME c_ts => pr_app is_closure thm vars fxy c_ts
- | NONE => brackify fxy
- [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
- | pr_term is_closure thm vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = Code_Thingol.unfold_abs t;
- fun pr ((v, pat), ty) =
- pr_bind is_closure thm NOBR ((SOME v, pat), ty)
- #>> (fn p => concat [str "fn", p, str "=>"]);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (ps @ [pr_term is_closure thm vars' NOBR t']) end
- | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) =
- (case Code_Thingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
- then pr_case is_closure thm vars fxy cases
- else pr_app is_closure thm vars fxy c_ts
- | NONE => pr_case is_closure thm vars fxy cases)
- and pr_app' is_closure thm vars (app as ((c, ((_, iss), tys)), ts)) =
- if is_cons c then
- let
- val k = length tys
- in if k < 2 then
- (str o deresolve) c :: map (pr_term is_closure thm vars BR) ts
- else if k = length ts then
- [(str o deresolve) c, Pretty.enum "," "(" ")" (map (pr_term is_closure thm vars NOBR) ts)]
- else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand k app)] end
- else if is_closure c
- then (str o deresolve) c @@ str "()"
- else
- (str o deresolve) c
- :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts
- and pr_app is_closure thm vars = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
- syntax_const thm vars
- and pr_bind' ((NONE, NONE), _) = str "_"
- | pr_bind' ((SOME v, NONE), _) = str v
- | pr_bind' ((NONE, SOME p), _) = p
- | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p]
- and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
- and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, body) = Code_Thingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
- |>> (fn p => semicolon [str "val", p, str "=", pr_term is_closure thm vars NOBR t])
- val (ps, vars') = fold_map pr binds vars;
- in
- Pretty.chunks [
- [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block,
- [str ("in"), Pretty.fbrk, pr_term is_closure thm vars' NOBR body] |> Pretty.block,
- str ("end")
- ]
- end
- | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) =
- let
- fun pr delim (pat, body) =
- let
- val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
- in
- concat [str delim, p, str "=>", pr_term is_closure thm vars' NOBR body]
- end;
- in
- brackets (
- str "case"
- :: pr_term is_closure thm vars NOBR t
- :: pr "of" clause
- :: map (pr "|") clauses
- )
- end
- | pr_case is_closure thm vars fxy ((_, []), _) =
- (concat o map str) ["raise", "Fail", "\"empty case\""];
- fun pr_stmt (MLExc (name, n)) =
- let
- val exc_str =
- (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
- in
- (concat o map str) (
- (if n = 0 then "val" else "fun")
- :: deresolve name
- :: replicate n "_"
- @ "="
- :: "raise"
- :: "Fail"
- @@ exc_str
- )
- end
- | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- (Code_Thingol.fold_constnames (insert (op =)) t []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts;
- in
- concat [
- str "val",
- (str o deresolve) name,
- str ":",
- pr_typ NOBR ty,
- str "=",
- pr_term (K false) thm vars NOBR t
- ]
- end
- | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) =
- let
- fun pr_funn definer (name, ((vs, ty), eqs as eq :: eqs')) =
- let
- val vs_dict = filter_out (null o snd) vs;
- val shift = if null eqs' then I else
- map (Pretty.block o single o Pretty.block o single);
- fun pr_eq definer ((ts, t), (thm, _)) =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts
- |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- concat (
- str definer
- :: (str o deresolve) name
- :: (if member (op =) pseudo_funs name then [str "()"]
- else pr_tyvar_dicts vs_dict
- @ map (pr_term (member (op =) pseudo_funs) thm vars BR) ts)
- @ str "="
- @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t
- )
- end
- in
- (Pretty.block o Pretty.fbreaks o shift) (
- pr_eq definer eq
- :: map (pr_eq "|") eqs'
- )
- end;
- fun pr_pseudo_fun name = concat [
- str "val",
- (str o deresolve) name,
- str "=",
- (str o deresolve) name,
- str "();"
- ];
- val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns);
- val pseudo_ps = map pr_pseudo_fun pseudo_funs;
- in Pretty.chunks (ps @ Pretty.block ([p, str ";"]) :: pseudo_ps) end
- | pr_stmt (MLDatas (datas as (data :: datas'))) =
- let
- fun pr_co (co, []) =
- str (deresolve co)
- | pr_co (co, tys) =
- concat [
- str (deresolve co),
- str "of",
- Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
- ];
- fun pr_data definer (tyco, (vs, [])) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- @@ str "EMPTY__"
- )
- | pr_data definer (tyco, (vs, cos)) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- :: separate (str "|") (map pr_co cos)
- );
- val (ps, p) = split_last
- (pr_data "datatype" data :: map (pr_data "and") datas');
- in Pretty.chunks (ps @| Pretty.block ([p, str ";"])) end
- | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
- let
- val w = Code_Printer.first_upper v ^ "_";
- fun pr_superclass_field (class, classrel) =
- (concat o map str) [
- deresolve classrel, ":", "'" ^ v, deresolve class
- ];
- fun pr_classparam_field (classparam, ty) =
- concat [
- (str o deresolve) classparam, str ":", pr_typ NOBR ty
- ];
- fun pr_classparam_proj (classparam, _) =
- semicolon [
- str "fun",
- (str o deresolve) classparam,
- Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)],
- str "=",
- str ("#" ^ deresolve classparam),
- str w
- ];
- fun pr_superclass_proj (_, classrel) =
- semicolon [
- str "fun",
- (str o deresolve) classrel,
- Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolve class)],
- str "=",
- str ("#" ^ deresolve classrel),
- str w
- ];
- in
- Pretty.chunks (
- concat [
- str ("type '" ^ v),
- (str o deresolve) class,
- str "=",
- Pretty.enum "," "{" "};" (
- map pr_superclass_field superclasses @ map pr_classparam_field classparams
- )
- ]
- :: map pr_superclass_proj superclasses
- @ map pr_classparam_proj classparams
- )
- end
- | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) =
- let
- fun pr_superclass (_, (classrel, dss)) =
- concat [
- (str o Long_Name.base_name o deresolve) classrel,
- str "=",
- pr_dicts NOBR [DictConst dss]
- ];
- fun pr_classparam ((classparam, c_inst), (thm, _)) =
- concat [
- (str o Long_Name.base_name o deresolve) classparam,
- str "=",
- pr_app (K false) thm reserved_names NOBR (c_inst, [])
- ];
- in
- semicolon ([
- str (if null arity then "val" else "fun"),
- (str o deresolve) inst ] @
- pr_tyvar_dicts arity @ [
- str "=",
- Pretty.enum "," "{" "}"
- (map pr_superclass superarities @ map pr_classparam classparam_insts),
- str ":",
- pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
- ])
- end;
- in pr_stmt end;
-
-fun pr_sml_module name content =
- Pretty.chunks (
- str ("structure " ^ name ^ " = ")
- :: str "struct"
- :: str ""
- :: content
- @ str ""
- @@ str ("end; (*struct " ^ name ^ "*)")
- );
-
-val literals_sml = Literals {
- literal_char = prefix "#" o quote o ML_Syntax.print_char,
- literal_string = quote o translate_string ML_Syntax.print_char,
- literal_numeral = fn unbounded => fn k =>
- if unbounded then "(" ^ string_of_int k ^ " : IntInf.int)"
- else string_of_int k,
- literal_list = Pretty.enum "," "[" "]",
- infix_cons = (7, "::")
-};
-
-
-(** OCaml serializer **)
-
-fun pr_ocaml_stmt labelled_name syntax_tyco syntax_const reserved_names deresolve is_cons =
- let
- fun pr_dicts fxy ds =
- let
- fun pr_dictvar (v, (_, 1)) = "_" ^ Code_Printer.first_upper v
- | pr_dictvar (v, (i, _)) = "_" ^ Code_Printer.first_upper v ^ string_of_int (i+1);
- fun pr_proj ps p =
- fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p
- fun pr_dict fxy (DictConst (inst, dss)) =
- brackify fxy ((str o deresolve) inst :: map (pr_dicts BR) dss)
- | pr_dict fxy (DictVar (classrels, v)) =
- pr_proj (map deresolve classrels) ((str o pr_dictvar) v)
- in case ds
- of [] => str "()"
- | [d] => pr_dict fxy d
- | _ :: _ => (Pretty.list "(" ")" o map (pr_dict NOBR)) ds
- end;
- fun pr_tyvar_dicts vs =
- vs
- |> map (fn (v, sort) => map_index (fn (i, _) =>
- DictVar ([], (v, (i, length sort)))) sort)
- |> map (pr_dicts BR);
- fun pr_tycoexpr fxy (tyco, tys) =
- let
- val tyco' = (str o deresolve) tyco
- in case map (pr_typ BR) tys
- of [] => tyco'
- | [p] => Pretty.block [p, Pretty.brk 1, tyco']
- | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
- end
- and pr_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
- of NONE => pr_tycoexpr fxy (tyco, tys)
- | SOME (i, pr) => pr pr_typ fxy tys)
- | pr_typ fxy (ITyVar v) = str ("'" ^ v);
- fun pr_term is_closure thm vars fxy (IConst c) =
- pr_app is_closure thm vars fxy (c, [])
- | pr_term is_closure thm vars fxy (IVar v) =
- str (Code_Printer.lookup_var vars v)
- | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
- (case Code_Thingol.unfold_const_app t
- of SOME c_ts => pr_app is_closure thm vars fxy c_ts
- | NONE =>
- brackify fxy [pr_term is_closure thm vars NOBR t1, pr_term is_closure thm vars BR t2])
- | pr_term is_closure thm vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = Code_Thingol.unfold_abs t;
- fun pr ((v, pat), ty) = pr_bind is_closure thm BR ((SOME v, pat), ty);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (str "fun" :: ps @ str "->" @@ pr_term is_closure thm vars' NOBR t') end
- | pr_term is_closure thm vars fxy (ICase (cases as (_, t0))) = (case Code_Thingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
- then pr_case is_closure thm vars fxy cases
- else pr_app is_closure thm vars fxy c_ts
- | NONE => pr_case is_closure thm vars fxy cases)
- and pr_app' is_closure thm vars (app as ((c, ((_, iss), tys)), ts)) =
- if is_cons c then
- if length tys = length ts
- then case ts
- of [] => [(str o deresolve) c]
- | [t] => [(str o deresolve) c, pr_term is_closure thm vars BR t]
- | _ => [(str o deresolve) c, Pretty.enum "," "(" ")"
- (map (pr_term is_closure thm vars NOBR) ts)]
- else [pr_term is_closure thm vars BR (Code_Thingol.eta_expand (length tys) app)]
- else if is_closure c
- then (str o deresolve) c @@ str "()"
- else (str o deresolve) c
- :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term is_closure thm vars BR) ts)
- and pr_app is_closure = gen_pr_app (pr_app' is_closure) (pr_term is_closure)
- syntax_const
- and pr_bind' ((NONE, NONE), _) = str "_"
- | pr_bind' ((SOME v, NONE), _) = str v
- | pr_bind' ((NONE, SOME p), _) = p
- | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v]
- and pr_bind is_closure = gen_pr_bind pr_bind' (pr_term is_closure)
- and pr_case is_closure thm vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, body) = Code_Thingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind is_closure thm NOBR ((NONE, SOME pat), ty)
- |>> (fn p => concat
- [str "let", p, str "=", pr_term is_closure thm vars NOBR t, str "in"])
- val (ps, vars') = fold_map pr binds vars;
- in
- brackify_block fxy (Pretty.chunks ps) []
- (pr_term is_closure thm vars' NOBR body)
- end
- | pr_case is_closure thm vars fxy (((t, ty), clause :: clauses), _) =
- let
- fun pr delim (pat, body) =
- let
- val (p, vars') = pr_bind is_closure thm NOBR ((NONE, SOME pat), ty) vars;
- in concat [str delim, p, str "->", pr_term is_closure thm vars' NOBR body] end;
- in
- brackets (
- str "match"
- :: pr_term is_closure thm vars NOBR t
- :: pr "with" clause
- :: map (pr "|") clauses
- )
- end
- | pr_case is_closure thm vars fxy ((_, []), _) =
- (concat o map str) ["failwith", "\"empty case\""];
- fun fish_params vars eqs =
- let
- fun fish_param _ (w as SOME _) = w
- | fish_param (IVar v) NONE = SOME v
- | fish_param _ NONE = NONE;
- fun fillup_param _ (_, SOME v) = v
- | fillup_param x (i, NONE) = x ^ string_of_int i;
- val fished1 = fold (map2 fish_param) eqs (replicate (length (hd eqs)) NONE);
- val x = Name.variant (map_filter I fished1) "x";
- val fished2 = map_index (fillup_param x) fished1;
- val (fished3, _) = Name.variants fished2 Name.context;
- val vars' = Code_Printer.intro_vars fished3 vars;
- in map (Code_Printer.lookup_var vars') fished3 end;
- fun pr_stmt (MLExc (name, n)) =
- let
- val exc_str =
- (ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
- in
- (concat o map str) (
- "let"
- :: deresolve name
- :: replicate n "_"
- @ "="
- :: "failwith"
- @@ exc_str
- )
- end
- | pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- (Code_Thingol.fold_constnames (insert (op =)) t []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts;
- in
- concat [
- str "let",
- (str o deresolve) name,
- str ":",
- pr_typ NOBR ty,
- str "=",
- pr_term (K false) thm vars NOBR t
- ]
- end
- | pr_stmt (MLFuns (funn :: funns, pseudo_funs)) =
- let
- fun pr_eq ((ts, t), (thm, _)) =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts
- |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in concat [
- (Pretty.block o Pretty.commas)
- (map (pr_term (member (op =) pseudo_funs) thm vars NOBR) ts),
- str "->",
- pr_term (member (op =) pseudo_funs) thm vars NOBR t
- ] end;
- fun pr_eqs is_pseudo [((ts, t), (thm, _))] =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts
- |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- concat (
- (if is_pseudo then [str "()"]
- else map (pr_term (member (op =) pseudo_funs) thm vars BR) ts)
- @ str "="
- @@ pr_term (member (op =) pseudo_funs) thm vars NOBR t
- )
- end
- | pr_eqs _ (eqs as (eq as (([_], _), _)) :: eqs') =
- Pretty.block (
- str "="
- :: Pretty.brk 1
- :: str "function"
- :: Pretty.brk 1
- :: pr_eq eq
- :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1]
- o single o pr_eq) eqs'
- )
- | pr_eqs _ (eqs as eq :: eqs') =
- let
- val consts = map_filter
- (fn c => if (is_some o syntax_const) c
- then NONE else (SOME o Long_Name.base_name o deresolve) c)
- ((fold o Code_Thingol.fold_constnames)
- (insert (op =)) (map (snd o fst) eqs) []);
- val vars = reserved_names
- |> Code_Printer.intro_vars consts;
- val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs;
- in
- Pretty.block (
- Pretty.breaks dummy_parms
- @ Pretty.brk 1
- :: str "="
- :: Pretty.brk 1
- :: str "match"
- :: Pretty.brk 1
- :: (Pretty.block o Pretty.commas) dummy_parms
- :: Pretty.brk 1
- :: str "with"
- :: Pretty.brk 1
- :: pr_eq eq
- :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1]
- o single o pr_eq) eqs'
- )
- end;
- fun pr_funn definer (name, ((vs, ty), eqs)) =
- concat (
- str definer
- :: (str o deresolve) name
- :: pr_tyvar_dicts (filter_out (null o snd) vs)
- @| pr_eqs (member (op =) pseudo_funs name) eqs
- );
- fun pr_pseudo_fun name = concat [
- str "let",
- (str o deresolve) name,
- str "=",
- (str o deresolve) name,
- str "();;"
- ];
- val (ps, p) = split_last (pr_funn "fun" funn :: map (pr_funn "and") funns);
- val (ps, p) = split_last
- (pr_funn "let rec" funn :: map (pr_funn "and") funns);
- val pseudo_ps = map pr_pseudo_fun pseudo_funs;
- in Pretty.chunks (ps @ Pretty.block ([p, str ";;"]) :: pseudo_ps) end
- | pr_stmt (MLDatas (datas as (data :: datas'))) =
- let
- fun pr_co (co, []) =
- str (deresolve co)
- | pr_co (co, tys) =
- concat [
- str (deresolve co),
- str "of",
- Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
- ];
- fun pr_data definer (tyco, (vs, [])) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- @@ str "EMPTY_"
- )
- | pr_data definer (tyco, (vs, cos)) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- :: separate (str "|") (map pr_co cos)
- );
- val (ps, p) = split_last
- (pr_data "type" data :: map (pr_data "and") datas');
- in Pretty.chunks (ps @| Pretty.block ([p, str ";;"])) end
- | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
- let
- val w = "_" ^ Code_Printer.first_upper v;
- fun pr_superclass_field (class, classrel) =
- (concat o map str) [
- deresolve classrel, ":", "'" ^ v, deresolve class
- ];
- fun pr_classparam_field (classparam, ty) =
- concat [
- (str o deresolve) classparam, str ":", pr_typ NOBR ty
- ];
- fun pr_classparam_proj (classparam, _) =
- concat [
- str "let",
- (str o deresolve) classparam,
- str w,
- str "=",
- str (w ^ "." ^ deresolve classparam ^ ";;")
- ];
- in Pretty.chunks (
- concat [
- str ("type '" ^ v),
- (str o deresolve) class,
- str "=",
- enum_default "unit;;" ";" "{" "};;" (
- map pr_superclass_field superclasses
- @ map pr_classparam_field classparams
- )
- ]
- :: map pr_classparam_proj classparams
- ) end
- | pr_stmt (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classparam_insts)))) =
- let
- fun pr_superclass (_, (classrel, dss)) =
- concat [
- (str o deresolve) classrel,
- str "=",
- pr_dicts NOBR [DictConst dss]
- ];
- fun pr_classparam_inst ((classparam, c_inst), (thm, _)) =
- concat [
- (str o deresolve) classparam,
- str "=",
- pr_app (K false) thm reserved_names NOBR (c_inst, [])
- ];
- in
- concat (
- str "let"
- :: (str o deresolve) inst
- :: pr_tyvar_dicts arity
- @ str "="
- @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [
- enum_default "()" ";" "{" "}" (map pr_superclass superarities
- @ map pr_classparam_inst classparam_insts),
- str ":",
- pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
- ]
- )
- end;
- in pr_stmt end;
-
-fun pr_ocaml_module name content =
- Pretty.chunks (
- str ("module " ^ name ^ " = ")
- :: str "struct"
- :: str ""
- :: content
- @ str ""
- @@ str ("end;; (*struct " ^ name ^ "*)")
- );
-
-val literals_ocaml = let
- fun chr i =
- let
- val xs = string_of_int i;
- val ys = replicate_string (3 - length (explode xs)) "0";
- in "\\" ^ ys ^ xs end;
- fun char_ocaml c =
- let
- val i = ord c;
- val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126
- then chr i else c
- in s end;
- fun bignum_ocaml k = if k <= 1073741823
- then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")"
- else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")"
-in Literals {
- literal_char = enclose "'" "'" o char_ocaml,
- literal_string = quote o translate_string char_ocaml,
- literal_numeral = fn unbounded => fn k => if k >= 0 then
- if unbounded then bignum_ocaml k
- else string_of_int k
- else
- if unbounded then "(Big_int.minus_big_int " ^ bignum_ocaml (~ k) ^ ")"
- else (enclose "(" ")" o prefix "-" o string_of_int o op ~) k,
- literal_list = Pretty.enum ";" "[" "]",
- infix_cons = (6, "::")
-} end;
-
-
-
-(** SML/OCaml generic part **)
-
-local
-
-datatype ml_node =
- Dummy of string
- | Stmt of string * ml_stmt
- | Module of string * ((Name.context * Name.context) * ml_node Graph.T);
-
-in
-
-fun ml_node_of_program labelled_name module_name reserved_names raw_module_alias program =
- let
- val module_alias = if is_some module_name then K module_name else raw_module_alias;
- val reserved_names = Name.make_context reserved_names;
- val empty_module = ((reserved_names, reserved_names), Graph.empty);
- fun map_node [] f = f
- | map_node (m::ms) f =
- Graph.default_node (m, Module (m, empty_module))
- #> Graph.map_node m (fn (Module (module_name, (nsp, nodes))) =>
- Module (module_name, (nsp, map_node ms f nodes)));
- fun map_nsp_yield [] f (nsp, nodes) =
- let
- val (x, nsp') = f nsp
- in (x, (nsp', nodes)) end
- | map_nsp_yield (m::ms) f (nsp, nodes) =
- let
- val (x, nodes') =
- nodes
- |> Graph.default_node (m, Module (m, empty_module))
- |> Graph.map_node_yield m (fn Module (d_module_name, nsp_nodes) =>
- let
- val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes
- in (x, Module (d_module_name, nsp_nodes')) end)
- in (x, (nsp, nodes')) end;
- fun map_nsp_fun_yield f (nsp_fun, nsp_typ) =
- let
- val (x, nsp_fun') = f nsp_fun
- in (x, (nsp_fun', nsp_typ)) end;
- fun map_nsp_typ_yield f (nsp_fun, nsp_typ) =
- let
- val (x, nsp_typ') = f nsp_typ
- in (x, (nsp_fun, nsp_typ')) end;
- val mk_name_module = Code_Printer.mk_name_module reserved_names NONE module_alias program;
- fun mk_name_stmt upper name nsp =
- let
- val (_, base) = Code_Printer.dest_name name;
- val base' = if upper then Code_Printer.first_upper base else base;
- val ([base''], nsp') = Name.variants [base'] nsp;
- in (base'', nsp') end;
- fun rearrange_fun name (tysm as (vs, ty), raw_eqs) =
- let
- val eqs = filter (snd o snd) raw_eqs;
- val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs
- of [(([], t), thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty
- then ([(([IVar "x"], t `$ IVar "x"), thm)], false)
- else (eqs, not (Code_Thingol.fold_constnames
- (fn name' => fn b => b orelse name = name') t false))
- | _ => (eqs, false)
- else (eqs, false)
- in ((name, (tysm, eqs')), is_value) end;
- fun check_kind [((name, (tysm, [(([], t), thm)])), true)] = MLVal (name, ((tysm, t), thm))
- | check_kind [((name, ((vs, ty), [])), _)] =
- MLExc (name, (length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty)
- | check_kind funns =
- MLFuns (map fst funns, map_filter
- (fn ((name, ((vs, _), [(([], _), _)])), _) =>
- if null (filter_out (null o snd) vs) then SOME name else NONE
- | _ => NONE) funns);
- fun add_funs stmts = fold_map
- (fn (name, Code_Thingol.Fun (_, stmt)) =>
- map_nsp_fun_yield (mk_name_stmt false name)
- #>> rpair (rearrange_fun name stmt)
- | (name, _) =>
- error ("Function block containing illegal statement: " ^ labelled_name name)
- ) stmts
- #>> (split_list #> apsnd check_kind);
- fun add_datatypes stmts =
- fold_map
- (fn (name, Code_Thingol.Datatype (_, stmt)) =>
- map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
- | (name, Code_Thingol.Datatypecons _) =>
- map_nsp_fun_yield (mk_name_stmt true name) #>> rpair NONE
- | (name, _) =>
- error ("Datatype block containing illegal statement: " ^ labelled_name name)
- ) stmts
- #>> (split_list #> apsnd (map_filter I
- #> (fn [] => error ("Datatype block without data statement: "
- ^ (commas o map (labelled_name o fst)) stmts)
- | stmts => MLDatas stmts)));
- fun add_class stmts =
- fold_map
- (fn (name, Code_Thingol.Class (_, stmt)) =>
- map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
- | (name, Code_Thingol.Classrel _) =>
- map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
- | (name, Code_Thingol.Classparam _) =>
- map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
- | (name, _) =>
- error ("Class block containing illegal statement: " ^ labelled_name name)
- ) stmts
- #>> (split_list #> apsnd (map_filter I
- #> (fn [] => error ("Class block without class statement: "
- ^ (commas o map (labelled_name o fst)) stmts)
- | [stmt] => MLClass stmt)));
- fun add_inst [(name, Code_Thingol.Classinst stmt)] =
- map_nsp_fun_yield (mk_name_stmt false name)
- #>> (fn base => ([base], MLClassinst (name, stmt)));
- fun add_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) =
- add_funs stmts
- | add_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) =
- add_datatypes stmts
- | add_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) =
- add_datatypes stmts
- | add_stmts ((stmts as (_, Code_Thingol.Class _)::_)) =
- add_class stmts
- | add_stmts ((stmts as (_, Code_Thingol.Classrel _)::_)) =
- add_class stmts
- | add_stmts ((stmts as (_, Code_Thingol.Classparam _)::_)) =
- add_class stmts
- | add_stmts ((stmts as [(_, Code_Thingol.Classinst _)])) =
- add_inst stmts
- | add_stmts stmts = error ("Illegal mutual dependencies: " ^
- (commas o map (labelled_name o fst)) stmts);
- fun add_stmts' stmts nsp_nodes =
- let
- val names as (name :: names') = map fst stmts;
- val deps =
- []
- |> fold (fold (insert (op =)) o Graph.imm_succs program) names
- |> subtract (op =) names;
- val (module_names, _) = (split_list o map Code_Printer.dest_name) names;
- val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
- handle Empty =>
- error ("Different namespace prefixes for mutual dependencies:\n"
- ^ commas (map labelled_name names)
- ^ "\n"
- ^ commas module_names);
- val module_name_path = Long_Name.explode module_name;
- fun add_dep name name' =
- let
- val module_name' = (mk_name_module o fst o Code_Printer.dest_name) name';
- in if module_name = module_name' then
- map_node module_name_path (Graph.add_edge (name, name'))
- else let
- val (common, (diff1 :: _, diff2 :: _)) = chop_prefix (op =)
- (module_name_path, Long_Name.explode module_name');
- in
- map_node common
- (fn node => Graph.add_edge_acyclic (diff1, diff2) node
- handle Graph.CYCLES _ => error ("Dependency "
- ^ quote name ^ " -> " ^ quote name'
- ^ " would result in module dependency cycle"))
- end end;
- in
- nsp_nodes
- |> map_nsp_yield module_name_path (add_stmts stmts)
- |-> (fn (base' :: bases', stmt') =>
- apsnd (map_node module_name_path (Graph.new_node (name, (Stmt (base', stmt')))
- #> fold2 (fn name' => fn base' =>
- Graph.new_node (name', (Dummy base'))) names' bases')))
- |> apsnd (fold (fn name => fold (add_dep name) deps) names)
- |> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names)
- end;
- val (_, nodes) = empty_module
- |> fold add_stmts' (map (AList.make (Graph.get_node program))
- (rev (Graph.strong_conn program)));
- fun deresolver prefix name =
- let
- val module_name = (fst o Code_Printer.dest_name) name;
- val module_name' = (Long_Name.explode o mk_name_module) module_name;
- val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name');
- val stmt_name =
- nodes
- |> fold (fn name => fn node => case Graph.get_node node name
- of Module (_, (_, node)) => node) module_name'
- |> (fn node => case Graph.get_node node name of Stmt (stmt_name, _) => stmt_name
- | Dummy stmt_name => stmt_name);
- in
- Long_Name.implode (remainder @ [stmt_name])
- end handle Graph.UNDEF _ =>
- error ("Unknown statement name: " ^ labelled_name name);
- in (deresolver, nodes) end;
-
-fun serialize_ml target compile pr_module pr_stmt raw_module_name labelled_name reserved_names includes raw_module_alias
- _ syntax_tyco syntax_const program stmt_names destination =
- let
- val is_cons = Code_Thingol.is_cons program;
- val present_stmt_names = Code_Target.stmt_names_of_destination destination;
- val is_present = not (null present_stmt_names);
- val module_name = if is_present then SOME "Code" else raw_module_name;
- val (deresolver, nodes) = ml_node_of_program labelled_name module_name
- reserved_names raw_module_alias program;
- val reserved_names = Code_Printer.make_vars reserved_names;
- fun pr_node prefix (Dummy _) =
- NONE
- | pr_node prefix (Stmt (_, stmt)) = if is_present andalso
- (null o filter (member (op =) present_stmt_names) o stmt_names_of) stmt
- then NONE
- else SOME
- (pr_stmt labelled_name syntax_tyco syntax_const reserved_names
- (deresolver prefix) is_cons stmt)
- | pr_node prefix (Module (module_name, (_, nodes))) =
- separate (str "")
- ((map_filter (pr_node (prefix @ [module_name]) o Graph.get_node nodes)
- o rev o flat o Graph.strong_conn) nodes)
- |> (if is_present then Pretty.chunks else pr_module module_name)
- |> SOME;
- val stmt_names' = (map o try)
- (deresolver (if is_some module_name then the_list module_name else [])) stmt_names;
- val p = Pretty.chunks (separate (str "") (map snd includes @ (map_filter
- (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes));
- in
- Code_Target.mk_serialization target
- (case compile of SOME compile => SOME (compile o Code_Target.code_of_pretty) | NONE => NONE)
- (fn NONE => Code_Target.code_writeln | SOME file => File.write file o Code_Target.code_of_pretty)
- (rpair stmt_names' o Code_Target.code_of_pretty) p destination
- end;
-
-end; (*local*)
-
-
-(** ML (system language) code for evaluation and instrumentalization **)
-
-fun eval_code_of some_target thy = Code_Target.serialize_custom thy (the_default target_Eval some_target,
- (fn _ => fn [] => serialize_ml target_SML (SOME (K ())) (K Pretty.chunks) pr_sml_stmt (SOME ""),
- literals_sml));
-
-
-(* evaluation *)
-
-fun eval some_target reff postproc thy t args =
- let
- val ctxt = ProofContext.init thy;
- fun evaluator naming program ((_, (_, ty)), t) deps =
- let
- val _ = if Code_Thingol.contains_dictvar t then
- error "Term to be evaluated contains free dictionaries" else ();
- val value_name = "Value.VALUE.value"
- val program' = program
- |> Graph.new_node (value_name,
- Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))])))
- |> fold (curry Graph.add_edge value_name) deps;
- val (value_code, [SOME value_name']) = eval_code_of some_target thy naming program' [value_name];
- val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name'
- ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
- in ML_Context.evaluate ctxt false reff sml_code end;
- in Code_Thingol.eval thy I postproc evaluator t end;
-
-
-(* instrumentalization by antiquotation *)
-
-local
-
-structure CodeAntiqData = ProofDataFun
-(
- type T = (string list * string list) * (bool * (string
- * (string * ((string * string) list * (string * string) list)) lazy));
- fun init _ = (([], []), (true, ("", Lazy.value ("", ([], [])))));
-);
-
-val is_first_occ = fst o snd o CodeAntiqData.get;
-
-fun delayed_code thy tycos consts () =
- let
- val (consts', (naming, program)) = Code_Thingol.consts_program thy consts;
- val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos;
- val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos');
- val (consts'', tycos'') = chop (length consts') target_names;
- val consts_map = map2 (fn const => fn NONE =>
- error ("Constant " ^ (quote o Code.string_of_const thy) const
- ^ "\nhas a user-defined serialization")
- | SOME const'' => (const, const'')) consts consts''
- val tycos_map = map2 (fn tyco => fn NONE =>
- error ("Type " ^ (quote o Sign.extern_type thy) tyco
- ^ "\nhas a user-defined serialization")
- | SOME tyco'' => (tyco, tyco'')) tycos tycos'';
- in (ml_code, (tycos_map, consts_map)) end;
-
-fun register_code new_tycos new_consts ctxt =
- let
- val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt;
- val tycos' = fold (insert (op =)) new_tycos tycos;
- val consts' = fold (insert (op =)) new_consts consts;
- val (struct_name', ctxt') = if struct_name = ""
- then ML_Antiquote.variant "Code" ctxt
- else (struct_name, ctxt);
- val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) tycos' consts');
- in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end;
-
-fun register_const const = register_code [] [const];
-
-fun register_datatype tyco constrs = register_code [tyco] constrs;
-
-fun print_const const all_struct_name tycos_map consts_map =
- (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const;
-
-fun print_datatype tyco constrs all_struct_name tycos_map consts_map =
- let
- val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode;
- fun check_base name name'' =
- if upperize (Long_Name.base_name name) = upperize name''
- then () else error ("Name as printed " ^ quote name''
- ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry.");
- val tyco'' = (the o AList.lookup (op =) tycos_map) tyco;
- val constrs'' = map (the o AList.lookup (op =) consts_map) constrs;
- val _ = check_base tyco tyco'';
- val _ = map2 check_base constrs constrs'';
- in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end;
-
-fun print_code struct_name is_first print_it ctxt =
- let
- val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt;
- val (raw_ml_code, (tycos_map, consts_map)) = Lazy.force acc_code;
- val ml_code = if is_first then "\nstructure " ^ struct_code_name
- ^ " =\nstruct\n\n" ^ raw_ml_code ^ "\nend;\n\n"
- else "";
- val all_struct_name = Long_Name.append struct_name struct_code_name;
- in (ml_code, print_it all_struct_name tycos_map consts_map) end;
-
-in
-
-fun ml_code_antiq raw_const {struct_name, background} =
- let
- val const = Code.check_const (ProofContext.theory_of background) raw_const;
- val is_first = is_first_occ background;
- val background' = register_const const background;
- in (print_code struct_name is_first (print_const const), background') end;
-
-fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} =
- let
- val thy = ProofContext.theory_of background;
- val tyco = Sign.intern_type thy raw_tyco;
- val constrs = map (Code.check_const thy) raw_constrs;
- val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
- val _ = if gen_eq_set (op =) (constrs, constrs') then ()
- else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
- val is_first = is_first_occ background;
- val background' = register_datatype tyco constrs background;
- in (print_code struct_name is_first (print_datatype tyco constrs), background') end;
-
-end; (*local*)
-
-
-(** Isar setup **)
-
-val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq);
-val _ = ML_Context.add_antiq "code_datatype" (fn _ =>
- (Args.tyname --| Scan.lift (Args.$$$ "=")
- -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term)))
- >> ml_code_datatype_antiq);
-
-fun isar_seri_sml module_name =
- Code_Target.parse_args (Scan.succeed ())
- #> (fn () => serialize_ml target_SML
- (SOME (use_text ML_Env.local_context (1, "generated code") false))
- pr_sml_module pr_sml_stmt module_name);
-
-fun isar_seri_ocaml module_name =
- Code_Target.parse_args (Scan.succeed ())
- #> (fn () => serialize_ml target_OCaml NONE
- pr_ocaml_module pr_ocaml_stmt module_name);
-
-val setup =
- Code_Target.add_target (target_SML, (isar_seri_sml, literals_sml))
- #> Code_Target.add_target (target_OCaml, (isar_seri_ocaml, literals_ocaml))
- #> Code_Target.extend_target (target_Eval, (target_SML, K I))
- #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> Code_Target.add_syntax_tyco target_OCaml "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names
- #> fold (Code_Target.add_reserved target_SML)
- ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)]
- #> fold (Code_Target.add_reserved target_OCaml) [
- "and", "as", "assert", "begin", "class",
- "constraint", "do", "done", "downto", "else", "end", "exception",
- "external", "false", "for", "fun", "function", "functor", "if",
- "in", "include", "inherit", "initializer", "lazy", "let", "match", "method",
- "module", "mutable", "new", "object", "of", "open", "or", "private", "rec",
- "sig", "struct", "then", "to", "true", "try", "type", "val",
- "virtual", "when", "while", "with"
- ]
- #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod"];
-
-end; (*struct*)
--- a/src/Tools/code/code_preproc.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,515 +0,0 @@
-(* Title: Tools/code/code_preproc.ML
- Author: Florian Haftmann, TU Muenchen
-
-Preprocessing code equations into a well-sorted system
-in a graph with explicit dependencies.
-*)
-
-signature CODE_PREPROC =
-sig
- val map_pre: (simpset -> simpset) -> theory -> theory
- val map_post: (simpset -> simpset) -> theory -> theory
- val add_inline: thm -> theory -> theory
- val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
- val del_functrans: string -> theory -> theory
- val simple_functrans: (theory -> thm list -> thm list option)
- -> theory -> (thm * bool) list -> (thm * bool) list option
- val print_codeproc: theory -> unit
-
- type code_algebra
- type code_graph
- val eqns: code_graph -> string -> (thm * bool) list
- val typ: code_graph -> string -> (string * sort) list * typ
- val all: code_graph -> string list
- val pretty: theory -> code_graph -> Pretty.T
- val obtain: theory -> string list -> term list -> code_algebra * code_graph
- val eval_conv: theory -> (sort -> sort)
- -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
- val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
- -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
-
- val setup: theory -> theory
-end
-
-structure Code_Preproc : CODE_PREPROC =
-struct
-
-(** preprocessor administration **)
-
-(* theory data *)
-
-datatype thmproc = Thmproc of {
- pre: simpset,
- post: simpset,
- functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
-};
-
-fun make_thmproc ((pre, post), functrans) =
- Thmproc { pre = pre, post = post, functrans = functrans };
-fun map_thmproc f (Thmproc { pre, post, functrans }) =
- make_thmproc (f ((pre, post), functrans));
-fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
- Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
- let
- val pre = Simplifier.merge_ss (pre1, pre2);
- val post = Simplifier.merge_ss (post1, post2);
- val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
- in make_thmproc ((pre, post), functrans) end;
-
-structure Code_Preproc_Data = TheoryDataFun
-(
- type T = thmproc;
- val empty = make_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
- fun copy spec = spec;
- val extend = copy;
- fun merge pp = merge_thmproc;
-);
-
-fun the_thmproc thy = case Code_Preproc_Data.get thy
- of Thmproc x => x;
-
-fun delete_force msg key xs =
- if AList.defined (op =) xs key then AList.delete (op =) key xs
- else error ("No such " ^ msg ^ ": " ^ quote key);
-
-fun map_data f thy =
- thy
- |> Code.purge_data
- |> (Code_Preproc_Data.map o map_thmproc) f;
-
-val map_pre = map_data o apfst o apfst;
-val map_post = map_data o apfst o apsnd;
-
-val add_inline = map_pre o MetaSimplifier.add_simp;
-val del_inline = map_pre o MetaSimplifier.del_simp;
-val add_post = map_post o MetaSimplifier.add_simp;
-val del_post = map_post o MetaSimplifier.del_simp;
-
-fun add_functrans (name, f) = (map_data o apsnd)
- (AList.update (op =) (name, (serial (), f)));
-
-fun del_functrans name = (map_data o apsnd)
- (delete_force "function transformer" name);
-
-
-(* post- and preprocessing *)
-
-fun apply_functrans thy c _ [] = []
- | apply_functrans thy c [] eqns = eqns
- | apply_functrans thy c functrans eqns = eqns
- |> perhaps (perhaps_loop (perhaps_apply functrans))
- |> Code.assert_eqns_const thy c;
-
-fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
-
-fun term_of_conv thy f =
- Thm.cterm_of thy
- #> f
- #> Thm.prop_of
- #> Logic.dest_equals
- #> snd;
-
-fun preprocess thy c eqns =
- let
- val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
- val functrans = (map (fn (_, (_, f)) => f thy) o #functrans
- o the_thmproc) thy;
- in
- eqns
- |> apply_functrans thy c functrans
- |> (map o apfst) (Code.rewrite_eqn pre)
- |> (map o apfst) (AxClass.unoverload thy)
- |> map (Code.assert_eqn thy)
- |> burrow_fst (Code.norm_args thy)
- |> burrow_fst (Code.norm_varnames thy)
- end;
-
-fun preprocess_conv thy ct =
- let
- val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
- in
- ct
- |> Simplifier.rewrite pre
- |> rhs_conv (AxClass.unoverload_conv thy)
- end;
-
-fun postprocess_conv thy ct =
- let
- val post = (Simplifier.theory_context thy o #post o the_thmproc) thy;
- in
- ct
- |> AxClass.overload_conv thy
- |> rhs_conv (Simplifier.rewrite post)
- end;
-
-fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
-
-fun print_codeproc thy =
- let
- val ctxt = ProofContext.init thy;
- val pre = (#pre o the_thmproc) thy;
- val post = (#post o the_thmproc) thy;
- val functrans = (map fst o #functrans o the_thmproc) thy;
- in
- (Pretty.writeln o Pretty.chunks) [
- Pretty.block [
- Pretty.str "preprocessing simpset:",
- Pretty.fbrk,
- Simplifier.pretty_ss ctxt pre
- ],
- Pretty.block [
- Pretty.str "postprocessing simpset:",
- Pretty.fbrk,
- Simplifier.pretty_ss ctxt post
- ],
- Pretty.block (
- Pretty.str "function transformers:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map Pretty.str) functrans
- )
- ]
- end;
-
-fun simple_functrans f thy eqns = case f thy (map fst eqns)
- of SOME thms' => SOME (map (rpair (forall snd eqns)) thms')
- | NONE => NONE;
-
-
-(** sort algebra and code equation graph types **)
-
-type code_algebra = (sort -> sort) * Sorts.algebra;
-type code_graph = (((string * sort) list * typ) * (thm * bool) list) Graph.T;
-
-fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr);
-fun typ eqngr = fst o Graph.get_node eqngr;
-fun all eqngr = Graph.keys eqngr;
-
-fun pretty thy eqngr =
- AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
- |> (map o apfst) (Code.string_of_const thy)
- |> sort (string_ord o pairself fst)
- |> map (fn (s, thms) =>
- (Pretty.block o Pretty.fbreaks) (
- Pretty.str s
- :: map (Display.pretty_thm o fst) thms
- ))
- |> Pretty.chunks;
-
-
-(** the Waisenhaus algorithm **)
-
-(* auxiliary *)
-
-fun is_proper_class thy = can (AxClass.get_info thy);
-
-fun complete_proper_sort thy =
- Sign.complete_sort thy #> filter (is_proper_class thy);
-
-fun inst_params thy tyco =
- map (fn (c, _) => AxClass.param_of_inst thy (c, tyco))
- o maps (#params o AxClass.get_info thy);
-
-fun consts_of thy eqns = [] |> (fold o fold o fold_aterms)
- (fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I)
- (map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns);
-
-fun tyscm_rhss_of thy c eqns =
- let
- val tyscm = case eqns of [] => Code.default_typscheme thy c
- | ((thm, _) :: _) => Code.typscheme_eqn thy thm;
- val rhss = consts_of thy eqns;
- in (tyscm, rhss) end;
-
-
-(* data structures *)
-
-datatype const = Fun of string | Inst of class * string;
-
-fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2)
- | const_ord (Inst class_tyco1, Inst class_tyco2) =
- prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2)
- | const_ord (Fun _, Inst _) = LESS
- | const_ord (Inst _, Fun _) = GREATER;
-
-type var = const * int;
-
-structure Vargraph =
- GraphFun(type key = var val ord = prod_ord const_ord int_ord);
-
-datatype styp = Tyco of string * styp list | Var of var | Free;
-
-fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys)
- | styp_of c_lhs (TFree (v, _)) = case c_lhs
- of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs)
- | NONE => Free;
-
-type vardeps_data = ((string * styp list) list * class list) Vargraph.T
- * (((string * sort) list * (thm * bool) list) Symtab.table
- * (class * string) list);
-
-val empty_vardeps_data : vardeps_data =
- (Vargraph.empty, (Symtab.empty, []));
-
-
-(* retrieving equations and instances from the background context *)
-
-fun obtain_eqns thy eqngr c =
- case try (Graph.get_node eqngr) c
- of SOME ((lhs, _), eqns) => ((lhs, []), [])
- | NONE => let
- val eqns = Code.these_eqns thy c
- |> preprocess thy c;
- val ((lhs, _), rhss) = tyscm_rhss_of thy c eqns;
- in ((lhs, rhss), eqns) end;
-
-fun obtain_instance thy arities (inst as (class, tyco)) =
- case AList.lookup (op =) arities inst
- of SOME classess => (classess, ([], []))
- | NONE => let
- val all_classes = complete_proper_sort thy [class];
- val superclasses = remove (op =) class all_classes
- val classess = map (complete_proper_sort thy)
- (Sign.arity_sorts thy tyco [class]);
- val inst_params = inst_params thy tyco all_classes;
- in (classess, (superclasses, inst_params)) end;
-
-
-(* computing instantiations *)
-
-fun add_classes thy arities eqngr c_k new_classes vardeps_data =
- let
- val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k;
- val diff_classes = new_classes |> subtract (op =) old_classes;
- in if null diff_classes then vardeps_data
- else let
- val c_ks = Vargraph.imm_succs (fst vardeps_data) c_k |> insert (op =) c_k;
- in
- vardeps_data
- |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes)
- |> fold (fn styp => fold (ensure_typmatch_inst thy arities eqngr styp) new_classes) styps
- |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
- end end
-and add_styp thy arities eqngr c_k tyco_styps vardeps_data =
- let
- val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k;
- in if member (op =) old_styps tyco_styps then vardeps_data
- else
- vardeps_data
- |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps)
- |> fold (ensure_typmatch_inst thy arities eqngr tyco_styps) classes
- end
-and add_dep thy arities eqngr c_k c_k' vardeps_data =
- let
- val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k;
- in
- vardeps_data
- |> add_classes thy arities eqngr c_k' classes
- |> apfst (Vargraph.add_edge (c_k, c_k'))
- end
-and ensure_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data =
- if can (Sign.arity_sorts thy tyco) [class]
- then vardeps_data
- |> ensure_inst thy arities eqngr (class, tyco)
- |> fold_index (fn (k, styp) =>
- ensure_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps
- else vardeps_data (*permissive!*)
-and ensure_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
- if member (op =) insts inst then vardeps_data
- else let
- val (classess, (superclasses, inst_params)) =
- obtain_instance thy arities inst;
- in
- vardeps_data
- |> (apsnd o apsnd) (insert (op =) inst)
- |> fold_index (fn (k, _) =>
- apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
- |> fold (fn superclass => ensure_inst thy arities eqngr (superclass, tyco)) superclasses
- |> fold (ensure_fun thy arities eqngr) inst_params
- |> fold_index (fn (k, classes) =>
- add_classes thy arities eqngr (Inst (class, tyco), k) classes
- #> fold (fn superclass =>
- add_dep thy arities eqngr (Inst (superclass, tyco), k)
- (Inst (class, tyco), k)) superclasses
- #> fold (fn inst_param =>
- add_dep thy arities eqngr (Fun inst_param, k)
- (Inst (class, tyco), k)
- ) inst_params
- ) classess
- end
-and ensure_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data =
- vardeps_data
- |> add_styp thy arities eqngr c_k tyco_styps
- | ensure_typmatch thy arities eqngr (Var c_k') c_k vardeps_data =
- vardeps_data
- |> add_dep thy arities eqngr c_k c_k'
- | ensure_typmatch thy arities eqngr Free c_k vardeps_data =
- vardeps_data
-and ensure_rhs thy arities eqngr (c', styps) vardeps_data =
- vardeps_data
- |> ensure_fun thy arities eqngr c'
- |> fold_index (fn (k, styp) =>
- ensure_typmatch thy arities eqngr styp (Fun c', k)) styps
-and ensure_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) =
- if Symtab.defined eqntab c then vardeps_data
- else let
- val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c;
- val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss;
- in
- vardeps_data
- |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns)))
- |> fold_index (fn (k, _) =>
- apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs
- |> fold_index (fn (k, (_, sort)) =>
- add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs
- |> fold (ensure_rhs thy arities eqngr) rhss'
- end;
-
-
-(* applying instantiations *)
-
-fun dicts_of thy (proj_sort, algebra) (T, sort) =
- let
- fun class_relation (x, _) _ = x;
- fun type_constructor tyco xs class =
- inst_params thy tyco (Sorts.complete_sort algebra [class])
- @ (maps o maps) fst xs;
- fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
- in
- flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra
- { class_relation = class_relation, type_constructor = type_constructor,
- type_variable = type_variable } (T, proj_sort sort)
- handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
- end;
-
-fun add_arity thy vardeps (class, tyco) =
- AList.default (op =)
- ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
- (0 upto Sign.arity_number thy tyco - 1));
-
-fun add_eqs thy vardeps (c, (proto_lhs, proto_eqns)) (rhss, eqngr) =
- if can (Graph.get_node eqngr) c then (rhss, eqngr)
- else let
- val lhs = map_index (fn (k, (v, _)) =>
- (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs;
- val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
- Vartab.update ((v, 0), sort)) lhs;
- val eqns = proto_eqns
- |> (map o apfst) (Code.inst_thm thy inst_tab);
- val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
- val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
- in (map (pair c) rhss' @ rhss, eqngr') end;
-
-fun extend_arities_eqngr thy cs ts (arities, eqngr) =
- let
- val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) =>
- insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts [];
- val (vardeps, (eqntab, insts)) = empty_vardeps_data
- |> fold (ensure_fun thy arities eqngr) cs
- |> fold (ensure_rhs thy arities eqngr) cs_rhss;
- val arities' = fold (add_arity thy vardeps) insts arities;
- val pp = Syntax.pp_global thy;
- val algebra = Sorts.subalgebra pp (is_proper_class thy)
- (AList.lookup (op =) arities') (Sign.classes_of thy);
- val (rhss, eqngr') = Symtab.fold (add_eqs thy vardeps) eqntab ([], eqngr);
- fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra)
- (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c);
- val eqngr'' = fold (fn (c, rhs) => fold
- (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr';
- in (algebra, (arities', eqngr'')) end;
-
-
-(** store for preprocessed arities and code equations **)
-
-structure Wellsorted = CodeDataFun
-(
- type T = ((string * class) * sort list) list * code_graph;
- val empty = ([], Graph.empty);
- fun purge thy cs (arities, eqngr) =
- let
- val del_cs = ((Graph.all_preds eqngr
- o filter (can (Graph.get_node eqngr))) cs);
- val del_arities = del_cs
- |> map_filter (AxClass.inst_of_param thy)
- |> maps (fn (c, tyco) =>
- (map (rpair tyco) o Sign.complete_sort thy o the_list
- o AxClass.class_of_param thy) c);
- val arities' = fold (AList.delete (op =)) del_arities arities;
- val eqngr' = Graph.del_nodes del_cs eqngr;
- in (arities', eqngr') end;
-);
-
-
-(** retrieval and evaluation interfaces **)
-
-fun obtain thy cs ts = apsnd snd
- (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
-
-fun prepare_sorts_typ prep_sort
- = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
-
-fun prepare_sorts prep_sort (Const (c, ty)) =
- Const (c, prepare_sorts_typ prep_sort ty)
- | prepare_sorts prep_sort (t1 $ t2) =
- prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
- | prepare_sorts prep_sort (Abs (v, ty, t)) =
- Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
- | prepare_sorts _ (t as Bound _) = t;
-
-fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct =
- let
- val pp = Syntax.pp_global thy;
- val ct = cterm_of proto_ct;
- val _ = (Sign.no_frees pp o map_types (K dummyT) o Sign.no_vars pp)
- (Thm.term_of ct);
- val thm = preprocess_conv thy ct;
- val ct' = Thm.rhs_of thm;
- val t' = Thm.term_of ct';
- val vs = Term.add_tfrees t' [];
- val consts = fold_aterms
- (fn Const (c, _) => insert (op =) c | _ => I) t' [];
-
- val t'' = prepare_sorts prep_sort t';
- val (algebra', eqngr') = obtain thy consts [t''];
- in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
-
-fun simple_evaluator evaluator algebra eqngr vs t ct =
- evaluator algebra eqngr vs t;
-
-fun eval_conv thy =
- let
- fun conclude_evaluation thm2 thm1 =
- let
- val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
- in
- Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
- error ("could not construct evaluation proof:\n"
- ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3])
- end;
- in gen_eval thy I conclude_evaluation end;
-
-fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy)
- (K o postproc (postprocess_term thy)) prep_sort (simple_evaluator evaluator);
-
-
-(** setup **)
-
-val setup =
- let
- fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
- fun add_del_attribute (name, (add, del)) =
- Code.add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
- || Scan.succeed (mk_attribute add))
- in
- add_del_attribute ("inline", (add_inline, del_inline))
- #> add_del_attribute ("post", (add_post, del_post))
- #> Code.add_attribute ("unfold", Scan.succeed (Thm.declaration_attribute
- (fn thm => Context.mapping (Codegen.add_unfold thm #> add_inline thm) I)))
- end;
-
-val _ =
- OuterSyntax.improper_command "print_codeproc" "print code preprocessor setup"
- OuterKeyword.diag (Scan.succeed
- (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
- (print_codeproc o Toplevel.theory_of)));
-
-end; (*struct*)
--- a/src/Tools/code/code_printer.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,318 +0,0 @@
-(* Title: Tools/code/code_printer.ML
- Author: Florian Haftmann, TU Muenchen
-
-Generic operations for pretty printing of target language code.
-*)
-
-signature CODE_PRINTER =
-sig
- val nerror: thm -> string -> 'a
-
- val @@ : 'a * 'a -> 'a list
- val @| : 'a list * 'a -> 'a list
- val str: string -> Pretty.T
- val concat: Pretty.T list -> Pretty.T
- val brackets: Pretty.T list -> Pretty.T
- val semicolon: Pretty.T list -> Pretty.T
- val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T
-
- val first_upper: string -> string
- val first_lower: string -> string
- type var_ctxt
- val make_vars: string list -> var_ctxt
- val intro_vars: string list -> var_ctxt -> var_ctxt
- val lookup_var: var_ctxt -> string -> string
-
- type literals
- val Literals: { literal_char: string -> string, literal_string: string -> string,
- literal_numeral: bool -> int -> string,
- literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }
- -> literals
- val literal_char: literals -> string -> string
- val literal_string: literals -> string -> string
- val literal_numeral: literals -> bool -> int -> string
- val literal_list: literals -> Pretty.T list -> Pretty.T
- val infix_cons: literals -> int * string
-
- type lrx
- val L: lrx
- val R: lrx
- val X: lrx
- type fixity
- val BR: fixity
- val NOBR: fixity
- val INFX: int * lrx -> fixity
- val APP: fixity
- val brackify: fixity -> Pretty.T list -> Pretty.T
- val brackify_infix: int * lrx -> fixity -> Pretty.T list -> Pretty.T
- val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T
-
- type itype = Code_Thingol.itype
- type iterm = Code_Thingol.iterm
- type const = Code_Thingol.const
- type dict = Code_Thingol.dict
- type tyco_syntax
- type const_syntax
- type proto_const_syntax
- val parse_infix: ('a -> 'b) -> lrx * int -> string
- -> int * ((fixity -> 'b -> Pretty.T)
- -> fixity -> 'a list -> Pretty.T)
- val parse_syntax: ('a -> 'b) -> OuterParse.token list
- -> (int * ((fixity -> 'b -> Pretty.T)
- -> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list
- val simple_const_syntax: (int * ((fixity -> iterm -> Pretty.T)
- -> fixity -> (iterm * itype) list -> Pretty.T)) option -> proto_const_syntax option
- val activate_const_syntax: theory -> literals
- -> proto_const_syntax -> Code_Thingol.naming -> const_syntax * Code_Thingol.naming
- val gen_pr_app: (thm -> var_ctxt -> const * iterm list -> Pretty.T list)
- -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
- -> (string -> const_syntax option)
- -> thm -> var_ctxt -> fixity -> const * iterm list -> Pretty.T
- val gen_pr_bind: ((string option * Pretty.T option) * itype -> Pretty.T)
- -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
- -> thm -> fixity
- -> (string option * iterm option) * itype -> var_ctxt -> Pretty.T * var_ctxt
-
- val mk_name_module: Name.context -> string option -> (string -> string option)
- -> 'a Graph.T -> string -> string
- val dest_name: string -> string * string
-end;
-
-structure Code_Printer : CODE_PRINTER =
-struct
-
-open Code_Thingol;
-
-fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm thm);
-
-(** assembling text pieces **)
-
-infixr 5 @@;
-infixr 5 @|;
-fun x @@ y = [x, y];
-fun xs @| y = xs @ [y];
-val str = PrintMode.setmp [] Pretty.str;
-val concat = Pretty.block o Pretty.breaks;
-val brackets = Pretty.enclose "(" ")" o Pretty.breaks;
-fun semicolon ps = Pretty.block [concat ps, str ";"];
-fun enum_default default sep opn cls [] = str default
- | enum_default default sep opn cls xs = Pretty.enum sep opn cls xs;
-
-
-(** names and variable name contexts **)
-
-type var_ctxt = string Symtab.table * Name.context;
-
-fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
- Name.make_context names);
-
-fun intro_vars names (namemap, namectxt) =
- let
- val (names', namectxt') = Name.variants names namectxt;
- val namemap' = fold2 (curry Symtab.update) names names' namemap;
- in (namemap', namectxt') end;
-
-fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
- of SOME name' => name'
- | NONE => error ("Invalid name in context: " ^ quote name);
-
-val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
-val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
-
-
-(** pretty literals **)
-
-datatype literals = Literals of {
- literal_char: string -> string,
- literal_string: string -> string,
- literal_numeral: bool -> int -> string,
- literal_list: Pretty.T list -> Pretty.T,
- infix_cons: int * string
-};
-
-fun dest_Literals (Literals lits) = lits;
-
-val literal_char = #literal_char o dest_Literals;
-val literal_string = #literal_string o dest_Literals;
-val literal_numeral = #literal_numeral o dest_Literals;
-val literal_list = #literal_list o dest_Literals;
-val infix_cons = #infix_cons o dest_Literals;
-
-
-(** syntax printer **)
-
-(* binding priorities *)
-
-datatype lrx = L | R | X;
-
-datatype fixity =
- BR
- | NOBR
- | INFX of (int * lrx);
-
-val APP = INFX (~1, L);
-
-fun fixity_lrx L L = false
- | fixity_lrx R R = false
- | fixity_lrx _ _ = true;
-
-fun fixity NOBR _ = false
- | fixity _ NOBR = false
- | fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) =
- pr < pr_ctxt
- orelse pr = pr_ctxt
- andalso fixity_lrx lr lr_ctxt
- orelse pr_ctxt = ~1
- | fixity BR (INFX _) = false
- | fixity _ _ = true;
-
-fun gen_brackify _ [p] = p
- | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
- | gen_brackify false (ps as _::_) = Pretty.block ps;
-
-fun brackify fxy_ctxt =
- gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks;
-
-fun brackify_infix infx fxy_ctxt =
- gen_brackify (fixity (INFX infx) fxy_ctxt) o Pretty.breaks;
-
-fun brackify_block fxy_ctxt p1 ps p2 =
- let val p = Pretty.block_enclose (p1, p2) ps
- in if fixity BR fxy_ctxt
- then Pretty.enclose "(" ")" [p]
- else p
- end;
-
-
-(* generic syntax *)
-
-type tyco_syntax = int * ((fixity -> itype -> Pretty.T)
- -> fixity -> itype list -> Pretty.T);
-type const_syntax = int * ((var_ctxt -> fixity -> iterm -> Pretty.T)
- -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
-type proto_const_syntax = int * (string list * (literals -> string list
- -> (var_ctxt -> fixity -> iterm -> Pretty.T)
- -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T));
-
-fun simple_const_syntax (SOME (n, f)) = SOME (n,
- ([], (fn _ => fn _ => fn pr => fn thm => fn vars => f (pr vars))))
- | simple_const_syntax NONE = NONE;
-
-fun activate_const_syntax thy literals (n, (cs, f)) naming =
- fold_map (Code_Thingol.ensure_declared_const thy) cs naming
- |-> (fn cs' => pair (n, f literals cs'));
-
-fun gen_pr_app pr_app pr_term syntax_const thm vars fxy (app as ((c, (_, tys)), ts)) =
- case syntax_const c
- of NONE => brackify fxy (pr_app thm vars app)
- | SOME (k, pr) =>
- let
- fun pr' fxy ts = pr (pr_term thm) thm vars fxy (ts ~~ curry Library.take k tys);
- in if k = length ts
- then pr' fxy ts
- else if k < length ts
- then case chop k ts of (ts1, ts2) =>
- brackify fxy (pr' APP ts1 :: map (pr_term thm vars BR) ts2)
- else pr_term thm vars fxy (Code_Thingol.eta_expand k app)
- end;
-
-fun gen_pr_bind pr_bind pr_term thm (fxy : fixity) ((v, pat), ty : itype) vars =
- let
- val vs = case pat
- of SOME pat => Code_Thingol.fold_varnames (insert (op =)) pat []
- | NONE => [];
- val vars' = intro_vars (the_list v) vars;
- val vars'' = intro_vars vs vars';
- val v' = Option.map (lookup_var vars') v;
- val pat' = Option.map (pr_term thm vars'' fxy) pat;
- in (pr_bind ((v', pat'), ty), vars'') end;
-
-
-(* mixfix syntax *)
-
-datatype 'a mixfix =
- Arg of fixity
- | Pretty of Pretty.T;
-
-fun mk_mixfix prep_arg (fixity_this, mfx) =
- let
- fun is_arg (Arg _) = true
- | is_arg _ = false;
- val i = (length o filter is_arg) mfx;
- fun fillin _ [] [] =
- []
- | fillin pr (Arg fxy :: mfx) (a :: args) =
- (pr fxy o prep_arg) a :: fillin pr mfx args
- | fillin pr (Pretty p :: mfx) args =
- p :: fillin pr mfx args;
- in
- (i, fn pr => fn fixity_ctxt => fn args =>
- gen_brackify (fixity fixity_this fixity_ctxt) (fillin pr mfx args))
- end;
-
-fun parse_infix prep_arg (x, i) s =
- let
- val l = case x of L => INFX (i, L) | _ => INFX (i, X);
- val r = case x of R => INFX (i, R) | _ => INFX (i, X);
- in
- mk_mixfix prep_arg (INFX (i, x),
- [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r])
- end;
-
-fun parse_mixfix prep_arg s =
- let
- val sym_any = Scan.one Symbol.is_regular;
- val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat (
- ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
- || ($$ "_" >> K (Arg BR))
- || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length))
- || (Scan.repeat1
- ( $$ "'" |-- sym_any
- || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
- sym_any) >> (Pretty o str o implode)));
- in case Scan.finite Symbol.stopper parse (Symbol.explode s)
- of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p)
- | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p)
- | _ => Scan.!!
- (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail ()
- end;
-
-val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr");
-
-fun parse_syntax prep_arg xs =
- Scan.option ((
- ((OuterParse.$$$ infixK >> K X)
- || (OuterParse.$$$ infixlK >> K L)
- || (OuterParse.$$$ infixrK >> K R))
- -- OuterParse.nat >> parse_infix prep_arg
- || Scan.succeed (parse_mixfix prep_arg))
- -- OuterParse.string
- >> (fn (parse, s) => parse s)) xs;
-
-val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK];
-
-
-(** module name spaces **)
-
-val dest_name =
- apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
-
-fun mk_name_module reserved_names module_prefix module_alias program =
- let
- fun mk_alias name = case module_alias name
- of SOME name' => name'
- | NONE => name
- |> Long_Name.explode
- |> map (fn name => (the_single o fst) (Name.variants [name] reserved_names))
- |> Long_Name.implode;
- fun mk_prefix name = case module_prefix
- of SOME module_prefix => Long_Name.append module_prefix name
- | NONE => name;
- val tab =
- Symtab.empty
- |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
- o fst o dest_name o fst)
- program
- in the o Symtab.lookup tab end;
-
-end; (*struct*)
--- a/src/Tools/code/code_target.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,629 +0,0 @@
-(* Title: Tools/code/code_target.ML
- Author: Florian Haftmann, TU Muenchen
-
-Serializer from intermediate language ("Thin-gol") to target languages.
-*)
-
-signature CODE_TARGET =
-sig
- include CODE_PRINTER
-
- type serializer
- val add_target: string * (serializer * literals) -> theory -> theory
- val extend_target: string *
- (string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program))
- -> theory -> theory
- val assert_target: theory -> string -> string
-
- type destination
- type serialization
- val parse_args: (OuterLex.token list -> 'a * OuterLex.token list)
- -> OuterLex.token list -> 'a
- val stmt_names_of_destination: destination -> string list
- val code_of_pretty: Pretty.T -> string
- val code_writeln: Pretty.T -> unit
- val mk_serialization: string -> ('a -> unit) option
- -> (Path.T option -> 'a -> unit)
- -> ('a -> string * string option list)
- -> 'a -> serialization
- val serialize: theory -> string -> string option -> OuterLex.token list
- -> Code_Thingol.naming -> Code_Thingol.program -> string list -> serialization
- val serialize_custom: theory -> string * (serializer * literals)
- -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
- val the_literals: theory -> string -> literals
- val compile: serialization -> unit
- val export: serialization -> unit
- val file: Path.T -> serialization -> unit
- val string: string list -> serialization -> string
- val code_of: theory -> string -> string
- -> string list -> (Code_Thingol.naming -> string list) -> string
- val shell_command: string (*theory name*) -> string (*export_code expr*) -> unit
- val code_width: int ref
-
- val allow_abort: string -> theory -> theory
- val add_syntax_class: string -> class -> string option -> theory -> theory
- val add_syntax_inst: string -> string * class -> bool -> theory -> theory
- val add_syntax_tyco: string -> string -> tyco_syntax option -> theory -> theory
- val add_syntax_const: string -> string -> proto_const_syntax option -> theory -> theory
- val add_reserved: string -> string -> theory -> theory
-end;
-
-structure Code_Target : CODE_TARGET =
-struct
-
-open Basic_Code_Thingol;
-open Code_Printer;
-
-(** basics **)
-
-datatype destination = Compile | Export | File of Path.T | String of string list;
-type serialization = destination -> (string * string option list) option;
-
-val code_width = ref 80; (*FIXME after Pretty module no longer depends on print mode*)
-fun code_setmp f = PrintMode.setmp [] (Pretty.setmp_margin (!code_width) f);
-fun code_of_pretty p = code_setmp Pretty.string_of p ^ "\n";
-fun code_writeln p = Pretty.setmp_margin (!code_width) Pretty.writeln p;
-
-(*FIXME why another code_setmp?*)
-fun compile f = (code_setmp f Compile; ());
-fun export f = (code_setmp f Export; ());
-fun file p f = (code_setmp f (File p); ());
-fun string stmts f = fst (the (code_setmp f (String stmts)));
-
-fun stmt_names_of_destination (String stmts) = stmts
- | stmt_names_of_destination _ = [];
-
-fun mk_serialization target (SOME comp) _ _ code Compile = (comp code; NONE)
- | mk_serialization target NONE _ _ _ Compile = error (target ^ ": no internal compilation")
- | mk_serialization target _ output _ code Export = (output NONE code ; NONE)
- | mk_serialization target _ output _ code (File file) = (output (SOME file) code; NONE)
- | mk_serialization target _ _ string code (String _) = SOME (string code);
-
-
-(** theory data **)
-
-datatype name_syntax_table = NameSyntaxTable of {
- class: string Symtab.table,
- instance: unit Symreltab.table,
- tyco: tyco_syntax Symtab.table,
- const: proto_const_syntax Symtab.table
-};
-
-fun mk_name_syntax_table ((class, instance), (tyco, const)) =
- NameSyntaxTable { class = class, instance = instance, tyco = tyco, const = const };
-fun map_name_syntax_table f (NameSyntaxTable { class, instance, tyco, const }) =
- mk_name_syntax_table (f ((class, instance), (tyco, const)));
-fun merge_name_syntax_table (NameSyntaxTable { class = class1, instance = instance1, tyco = tyco1, const = const1 },
- NameSyntaxTable { class = class2, instance = instance2, tyco = tyco2, const = const2 }) =
- mk_name_syntax_table (
- (Symtab.join (K snd) (class1, class2),
- Symreltab.join (K snd) (instance1, instance2)),
- (Symtab.join (K snd) (tyco1, tyco2),
- Symtab.join (K snd) (const1, const2))
- );
-
-type serializer =
- string option (*module name*)
- -> OuterLex.token list (*arguments*)
- -> (string -> string) (*labelled_name*)
- -> string list (*reserved symbols*)
- -> (string * Pretty.T) list (*includes*)
- -> (string -> string option) (*module aliasses*)
- -> (string -> string option) (*class syntax*)
- -> (string -> tyco_syntax option)
- -> (string -> const_syntax option)
- -> Code_Thingol.program
- -> string list (*selected statements*)
- -> serialization;
-
-datatype serializer_entry = Serializer of serializer * literals
- | Extends of string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program);
-
-datatype target = Target of {
- serial: serial,
- serializer: serializer_entry,
- reserved: string list,
- includes: (Pretty.T * string list) Symtab.table,
- name_syntax_table: name_syntax_table,
- module_alias: string Symtab.table
-};
-
-fun make_target ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))) =
- Target { serial = serial, serializer = serializer, reserved = reserved,
- includes = includes, name_syntax_table = name_syntax_table, module_alias = module_alias };
-fun map_target f ( Target { serial, serializer, reserved, includes, name_syntax_table, module_alias } ) =
- make_target (f ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))));
-fun merge_target strict target (Target { serial = serial1, serializer = serializer,
- reserved = reserved1, includes = includes1,
- name_syntax_table = name_syntax_table1, module_alias = module_alias1 },
- Target { serial = serial2, serializer = _,
- reserved = reserved2, includes = includes2,
- name_syntax_table = name_syntax_table2, module_alias = module_alias2 }) =
- if serial1 = serial2 orelse not strict then
- make_target ((serial1, serializer),
- ((merge (op =) (reserved1, reserved2), Symtab.merge (op =) (includes1, includes2)),
- (merge_name_syntax_table (name_syntax_table1, name_syntax_table2),
- Symtab.join (K snd) (module_alias1, module_alias2))
- ))
- else
- error ("Incompatible serializers: " ^ quote target);
-
-structure CodeTargetData = TheoryDataFun
-(
- type T = target Symtab.table * string list;
- val empty = (Symtab.empty, []);
- val copy = I;
- val extend = I;
- fun merge _ ((target1, exc1) : T, (target2, exc2)) =
- (Symtab.join (merge_target true) (target1, target2), Library.merge (op =) (exc1, exc2));
-);
-
-fun the_serializer (Target { serializer, ... }) = serializer;
-fun the_reserved (Target { reserved, ... }) = reserved;
-fun the_includes (Target { includes, ... }) = includes;
-fun the_name_syntax (Target { name_syntax_table = NameSyntaxTable x, ... }) = x;
-fun the_module_alias (Target { module_alias , ... }) = module_alias;
-
-val abort_allowed = snd o CodeTargetData.get;
-
-fun assert_target thy target =
- case Symtab.lookup (fst (CodeTargetData.get thy)) target
- of SOME data => target
- | NONE => error ("Unknown code target language: " ^ quote target);
-
-fun put_target (target, seri) thy =
- let
- val lookup_target = Symtab.lookup (fst (CodeTargetData.get thy));
- val _ = case seri
- of Extends (super, _) => if is_some (lookup_target super) then ()
- else error ("Unknown code target language: " ^ quote super)
- | _ => ();
- val overwriting = case (Option.map the_serializer o lookup_target) target
- of NONE => false
- | SOME (Extends _) => true
- | SOME (Serializer _) => (case seri
- of Extends _ => error ("Will not overwrite existing target " ^ quote target)
- | _ => true);
- val _ = if overwriting
- then warning ("Overwriting existing target " ^ quote target)
- else ();
- in
- thy
- |> (CodeTargetData.map o apfst oo Symtab.map_default)
- (target, make_target ((serial (), seri), (([], Symtab.empty),
- (mk_name_syntax_table ((Symtab.empty, Symreltab.empty), (Symtab.empty, Symtab.empty)),
- Symtab.empty))))
- ((map_target o apfst o apsnd o K) seri)
- end;
-
-fun add_target (target, seri) = put_target (target, Serializer seri);
-fun extend_target (target, (super, modify)) =
- put_target (target, Extends (super, modify));
-
-fun map_target_data target f thy =
- let
- val _ = assert_target thy target;
- in
- thy
- |> (CodeTargetData.map o apfst o Symtab.map_entry target o map_target) f
- end;
-
-fun map_reserved target =
- map_target_data target o apsnd o apfst o apfst;
-fun map_includes target =
- map_target_data target o apsnd o apfst o apsnd;
-fun map_name_syntax target =
- map_target_data target o apsnd o apsnd o apfst o map_name_syntax_table;
-fun map_module_alias target =
- map_target_data target o apsnd o apsnd o apsnd;
-
-
-(** serializer configuration **)
-
-(* data access *)
-
-local
-
-fun cert_class thy class =
- let
- val _ = AxClass.get_info thy class;
- in class end;
-
-fun read_class thy = cert_class thy o Sign.intern_class thy;
-
-fun cert_tyco thy tyco =
- let
- val _ = if Sign.declared_tyname thy tyco then ()
- else error ("No such type constructor: " ^ quote tyco);
- in tyco end;
-
-fun read_tyco thy = cert_tyco thy o Sign.intern_type thy;
-
-fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy =
- let
- val class = prep_class thy raw_class;
- in case raw_syn
- of SOME syntax =>
- thy
- |> (map_name_syntax target o apfst o apfst)
- (Symtab.update (class, syntax))
- | NONE =>
- thy
- |> (map_name_syntax target o apfst o apfst)
- (Symtab.delete_safe class)
- end;
-
-fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy =
- let
- val inst = (prep_class thy raw_class, prep_tyco thy raw_tyco);
- in if add_del then
- thy
- |> (map_name_syntax target o apfst o apsnd)
- (Symreltab.update (inst, ()))
- else
- thy
- |> (map_name_syntax target o apfst o apsnd)
- (Symreltab.delete_safe inst)
- end;
-
-fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy =
- let
- val tyco = prep_tyco thy raw_tyco;
- fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco
- then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
- else syntax
- in case raw_syn
- of SOME syntax =>
- thy
- |> (map_name_syntax target o apsnd o apfst)
- (Symtab.update (tyco, check_args syntax))
- | NONE =>
- thy
- |> (map_name_syntax target o apsnd o apfst)
- (Symtab.delete_safe tyco)
- end;
-
-fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
- let
- val c = prep_const thy raw_c;
- fun check_args (syntax as (n, _)) = if n > Code.no_args thy c
- then error ("Too many arguments in syntax for constant " ^ quote c)
- else syntax;
- in case raw_syn
- of SOME syntax =>
- thy
- |> (map_name_syntax target o apsnd o apsnd)
- (Symtab.update (c, check_args syntax))
- | NONE =>
- thy
- |> (map_name_syntax target o apsnd o apsnd)
- (Symtab.delete_safe c)
- end;
-
-fun add_reserved target =
- let
- fun add sym syms = if member (op =) syms sym
- then error ("Reserved symbol " ^ quote sym ^ " already declared")
- else insert (op =) sym syms
- in map_reserved target o add end;
-
-fun gen_add_include read_const target args thy =
- let
- fun add (name, SOME (content, raw_cs)) incls =
- let
- val _ = if Symtab.defined incls name
- then warning ("Overwriting existing include " ^ name)
- else ();
- val cs = map (read_const thy) raw_cs;
- in Symtab.update (name, (str content, cs)) incls end
- | add (name, NONE) incls = Symtab.delete name incls;
- in map_includes target (add args) thy end;
-
-val add_include = gen_add_include Code.check_const;
-val add_include_cmd = gen_add_include Code.read_const;
-
-fun add_module_alias target (thyname, modlname) =
- let
- val xs = Long_Name.explode modlname;
- val xs' = map (Name.desymbolize true) xs;
- in if xs' = xs
- then map_module_alias target (Symtab.update (thyname, modlname))
- else error ("Invalid module name: " ^ quote modlname ^ "\n"
- ^ "perhaps try " ^ quote (Long_Name.implode xs'))
- end;
-
-fun gen_allow_abort prep_const raw_c thy =
- let
- val c = prep_const thy raw_c;
- in thy |> (CodeTargetData.map o apsnd) (insert (op =) c) end;
-
-fun zip_list (x::xs) f g =
- f
- #-> (fn y =>
- fold_map (fn x => g |-- f >> pair x) xs
- #-> (fn xys => pair ((x, y) :: xys)));
-
-
-(* concrete syntax *)
-
-structure P = OuterParse
-and K = OuterKeyword
-
-fun parse_multi_syntax parse_thing parse_syntax =
- P.and_list1 parse_thing
- #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
- (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")"));
-
-in
-
-val add_syntax_class = gen_add_syntax_class cert_class (K I);
-val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco;
-val add_syntax_tyco = gen_add_syntax_tyco cert_tyco;
-val add_syntax_const = gen_add_syntax_const (K I);
-val allow_abort = gen_allow_abort (K I);
-val add_reserved = add_reserved;
-
-val add_syntax_class_cmd = gen_add_syntax_class read_class Code.read_const;
-val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
-val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
-val add_syntax_const_cmd = gen_add_syntax_const Code.read_const;
-val allow_abort_cmd = gen_allow_abort Code.read_const;
-
-fun the_literals thy =
- let
- val (targets, _) = CodeTargetData.get thy;
- fun literals target = case Symtab.lookup targets target
- of SOME data => (case the_serializer data
- of Serializer (_, literals) => literals
- | Extends (super, _) => literals super)
- | NONE => error ("Unknown code target language: " ^ quote target);
- in literals end;
-
-
-(** serializer usage **)
-
-(* montage *)
-
-local
-
-fun labelled_name thy program name = case Graph.get_node program name
- of Code_Thingol.Fun (c, _) => quote (Code.string_of_const thy c)
- | Code_Thingol.Datatype (tyco, _) => "type " ^ quote (Sign.extern_type thy tyco)
- | Code_Thingol.Datatypecons (c, _) => quote (Code.string_of_const thy c)
- | Code_Thingol.Class (class, _) => "class " ^ quote (Sign.extern_class thy class)
- | Code_Thingol.Classrel (sub, super) => let
- val Code_Thingol.Class (sub, _) = Graph.get_node program sub
- val Code_Thingol.Class (super, _) = Graph.get_node program super
- in quote (Sign.extern_class thy sub ^ " < " ^ Sign.extern_class thy super) end
- | Code_Thingol.Classparam (c, _) => quote (Code.string_of_const thy c)
- | Code_Thingol.Classinst ((class, (tyco, _)), _) => let
- val Code_Thingol.Class (class, _) = Graph.get_node program class
- val Code_Thingol.Datatype (tyco, _) = Graph.get_node program tyco
- in quote (Sign.extern_type thy tyco ^ " :: " ^ Sign.extern_class thy class) end
-
-fun activate_syntax lookup_name src_tab = Symtab.empty
- |> fold_map (fn thing_identifier => fn tab => case lookup_name thing_identifier
- of SOME name => (SOME name,
- Symtab.update_new (name, the (Symtab.lookup src_tab thing_identifier)) tab)
- | NONE => (NONE, tab)) (Symtab.keys src_tab)
- |>> map_filter I;
-
-fun activate_const_syntax thy literals src_tab naming = (Symtab.empty, naming)
- |> fold_map (fn thing_identifier => fn (tab, naming) =>
- case Code_Thingol.lookup_const naming thing_identifier
- of SOME name => let
- val (syn, naming') = Code_Printer.activate_const_syntax thy
- literals (the (Symtab.lookup src_tab thing_identifier)) naming
- in (SOME name, (Symtab.update_new (name, syn) tab, naming')) end
- | NONE => (NONE, (tab, naming))) (Symtab.keys src_tab)
- |>> map_filter I;
-
-fun invoke_serializer thy abortable serializer literals reserved abs_includes
- module_alias class instance tyco const module args naming program2 names1 =
- let
- val (names_class, class') =
- activate_syntax (Code_Thingol.lookup_class naming) class;
- val names_inst = map_filter (Code_Thingol.lookup_instance naming)
- (Symreltab.keys instance);
- val (names_tyco, tyco') =
- activate_syntax (Code_Thingol.lookup_tyco naming) tyco;
- val (names_const, (const', _)) =
- activate_const_syntax thy literals const naming;
- val names_hidden = names_class @ names_inst @ names_tyco @ names_const;
- val names2 = subtract (op =) names_hidden names1;
- val program3 = Graph.subgraph (not o member (op =) names_hidden) program2;
- val names_all = Graph.all_succs program3 names2;
- val includes = abs_includes names_all;
- val program4 = Graph.subgraph (member (op =) names_all) program3;
- val empty_funs = filter_out (member (op =) abortable)
- (Code_Thingol.empty_funs program3);
- val _ = if null empty_funs then () else error ("No code equations for "
- ^ commas (map (Sign.extern_const thy) empty_funs));
- in
- serializer module args (labelled_name thy program2) reserved includes
- (Symtab.lookup module_alias) (Symtab.lookup class')
- (Symtab.lookup tyco') (Symtab.lookup const')
- program4 names2
- end;
-
-fun mount_serializer thy alt_serializer target module args naming program names =
- let
- val (targets, abortable) = CodeTargetData.get thy;
- fun collapse_hierarchy target =
- let
- val data = case Symtab.lookup targets target
- of SOME data => data
- | NONE => error ("Unknown code target language: " ^ quote target);
- in case the_serializer data
- of Serializer _ => (I, data)
- | Extends (super, modify) => let
- val (modify', data') = collapse_hierarchy super
- in (modify' #> modify naming, merge_target false target (data', data)) end
- end;
- val (modify, data) = collapse_hierarchy target;
- val (serializer, _) = the_default (case the_serializer data
- of Serializer seri => seri) alt_serializer;
- val reserved = the_reserved data;
- fun select_include names_all (name, (content, cs)) =
- if null cs then SOME (name, content)
- else if exists (fn c => case Code_Thingol.lookup_const naming c
- of SOME name => member (op =) names_all name
- | NONE => false) cs
- then SOME (name, content) else NONE;
- fun includes names_all = map_filter (select_include names_all)
- ((Symtab.dest o the_includes) data);
- val module_alias = the_module_alias data;
- val { class, instance, tyco, const } = the_name_syntax data;
- val literals = the_literals thy target;
- in
- invoke_serializer thy abortable serializer literals reserved
- includes module_alias class instance tyco const module args naming (modify program) names
- end;
-
-in
-
-fun serialize thy = mount_serializer thy NONE;
-
-fun serialize_custom thy (target_name, seri) naming program names =
- mount_serializer thy (SOME seri) target_name NONE [] naming program names (String [])
- |> the;
-
-end; (* local *)
-
-fun parse_args f args =
- case Scan.read OuterLex.stopper f args
- of SOME x => x
- | NONE => error "Bad serializer arguments";
-
-
-(* code presentation *)
-
-fun code_of thy target module_name cs names_stmt =
- let
- val (names_cs, (naming, program)) = Code_Thingol.consts_program thy cs;
- in
- string (names_stmt naming) (serialize thy target (SOME module_name) []
- naming program names_cs)
- end;
-
-
-(* code generation *)
-
-fun transitivly_non_empty_funs thy naming program =
- let
- val cs = subtract (op =) (abort_allowed thy) (Code_Thingol.empty_funs program);
- val names = map_filter (Code_Thingol.lookup_const naming) cs;
- in subtract (op =) (Graph.all_preds program names) (Graph.keys program) end;
-
-fun read_const_exprs thy cs =
- let
- val (cs1, cs2) = Code_Thingol.read_const_exprs thy cs;
- val (names3, (naming, program)) = Code_Thingol.consts_program thy cs2;
- val names4 = transitivly_non_empty_funs thy naming program;
- val cs5 = map_filter
- (fn (c, name) => if member (op =) names4 name then SOME c else NONE) (cs2 ~~ names3);
- in fold (insert (op =)) cs5 cs1 end;
-
-fun cached_program thy =
- let
- val (naming, program) = Code_Thingol.cached_program thy;
- in (transitivly_non_empty_funs thy naming program, (naming, program)) end
-
-fun export_code thy cs seris =
- let
- val (cs', (naming, program)) = if null cs then cached_program thy
- else Code_Thingol.consts_program thy cs;
- fun mk_seri_dest dest = case dest
- of NONE => compile
- | SOME "-" => export
- | SOME f => file (Path.explode f)
- val _ = map (fn (((target, module), dest), args) =>
- (mk_seri_dest dest (serialize thy target module args naming program cs'))) seris;
- in () end;
-
-fun export_code_cmd raw_cs seris thy = export_code thy (read_const_exprs thy raw_cs) seris;
-
-
-(** Isar setup **)
-
-val (inK, module_nameK, fileK) = ("in", "module_name", "file");
-
-val code_exprP =
- (Scan.repeat P.term_group
- -- Scan.repeat (P.$$$ inK |-- P.name
- -- Scan.option (P.$$$ module_nameK |-- P.name)
- -- Scan.option (P.$$$ fileK |-- P.name)
- -- Scan.optional (P.$$$ "(" |-- Args.parse --| P.$$$ ")") []
- ) >> (fn (raw_cs, seris) => export_code_cmd raw_cs seris));
-
-val _ = List.app OuterKeyword.keyword [inK, module_nameK, fileK];
-
-val _ =
- OuterSyntax.command "code_class" "define code syntax for class" K.thy_decl (
- parse_multi_syntax P.xname (Scan.option P.string)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns)
- );
-
-val _ =
- OuterSyntax.command "code_instance" "define code syntax for instance" K.thy_decl (
- parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname)
- ((P.minus >> K true) || Scan.succeed false)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns)
- );
-
-val _ =
- OuterSyntax.command "code_type" "define code syntax for type constructor" K.thy_decl (
- parse_multi_syntax P.xname (parse_syntax I)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns)
- );
-
-val _ =
- OuterSyntax.command "code_const" "define code syntax for constant" K.thy_decl (
- parse_multi_syntax P.term_group (parse_syntax fst)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_const, syn) => add_syntax_const_cmd target raw_const
- (Code_Printer.simple_const_syntax syn)) syns)
- );
-
-val _ =
- OuterSyntax.command "code_reserved" "declare words as reserved for target language" K.thy_decl (
- P.name -- Scan.repeat1 P.name
- >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)
- );
-
-val _ =
- OuterSyntax.command "code_include" "declare piece of code to be included in generated code" K.thy_decl (
- P.name -- P.name -- (P.text :|-- (fn "-" => Scan.succeed NONE
- | s => Scan.optional (P.$$$ "attach" |-- Scan.repeat1 P.term) [] >> pair s >> SOME))
- >> (fn ((target, name), content_consts) =>
- (Toplevel.theory o add_include_cmd target) (name, content_consts))
- );
-
-val _ =
- OuterSyntax.command "code_modulename" "alias module to other name" K.thy_decl (
- P.name -- Scan.repeat1 (P.name -- P.name)
- >> (fn (target, modlnames) => (Toplevel.theory o fold (add_module_alias target)) modlnames)
- );
-
-val _ =
- OuterSyntax.command "code_abort" "permit constant to be implemented as program abort" K.thy_decl (
- Scan.repeat1 P.term_group >> (Toplevel.theory o fold allow_abort_cmd)
- );
-
-val _ =
- OuterSyntax.command "export_code" "generate executable code for constants"
- K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
-
-fun shell_command thyname cmd = Toplevel.program (fn _ =>
- (use_thy thyname; case Scan.read OuterLex.stopper (P.!!! code_exprP)
- ((filter OuterLex.is_proper o OuterSyntax.scan Position.none) cmd)
- of SOME f => (writeln "Now generating code..."; f (theory thyname))
- | NONE => error ("Bad directive " ^ quote cmd)))
- handle TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure;
-
-end; (*local*)
-
-end; (*struct*)
--- a/src/Tools/code/code_thingol.ML Thu Jun 18 17:14:08 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,876 +0,0 @@
-(* Title: Tools/code/code_thingol.ML
- Author: Florian Haftmann, TU Muenchen
-
-Intermediate language ("Thin-gol") representing executable code.
-Representation and translation.
-*)
-
-infix 8 `%%;
-infix 4 `$;
-infix 4 `$$;
-infixr 3 `|->;
-infixr 3 `|-->;
-
-signature BASIC_CODE_THINGOL =
-sig
- type vname = string;
- datatype dict =
- DictConst of string * dict list list
- | DictVar of string list * (vname * (int * int));
- datatype itype =
- `%% of string * itype list
- | ITyVar of vname;
- type const = string * ((itype list * dict list list) * itype list (*types of arguments*))
- datatype iterm =
- IConst of const
- | IVar of vname
- | `$ of iterm * iterm
- | `|-> of (vname * itype) * iterm
- | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
- (*((term, type), [(selector pattern, body term )]), primitive term)*)
- val `$$ : iterm * iterm list -> iterm;
- val `|--> : (vname * itype) list * iterm -> iterm;
- type typscheme = (vname * sort) list * itype;
-end;
-
-signature CODE_THINGOL =
-sig
- include BASIC_CODE_THINGOL
- val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list
- val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
- val unfold_fun: itype -> itype list * itype
- val unfold_app: iterm -> iterm * iterm list
- val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option
- val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm
- val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option
- val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm
- val unfold_const_app: iterm -> (const * iterm list) option
- val collapse_let: ((vname * itype) * iterm) * iterm
- -> (iterm * itype) * (iterm * iterm) list
- val eta_expand: int -> const * iterm list -> iterm
- val contains_dictvar: iterm -> bool
- val locally_monomorphic: iterm -> bool
- val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
- val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
- val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a
-
- type naming
- val empty_naming: naming
- val lookup_class: naming -> class -> string option
- val lookup_classrel: naming -> class * class -> string option
- val lookup_tyco: naming -> string -> string option
- val lookup_instance: naming -> class * string -> string option
- val lookup_const: naming -> string -> string option
- val ensure_declared_const: theory -> string -> naming -> string * naming
-
- datatype stmt =
- NoStmt
- | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)
- | Datatype of string * ((vname * sort) list * (string * itype list) list)
- | Datatypecons of string * string
- | Class of class * (vname * ((class * string) list * (string * itype) list))
- | Classrel of class * class
- | Classparam of string * class
- | Classinst of (class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list)
- type program = stmt Graph.T
- val empty_funs: program -> string list
- val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm
- val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt
- val is_cons: program -> string -> bool
- val contr_classparam_typs: program -> string -> itype option list
-
- val read_const_exprs: theory -> string list -> string list * string list
- val consts_program: theory -> string list -> string list * (naming * program)
- val cached_program: theory -> naming * program
- val eval_conv: theory -> (sort -> sort)
- -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm)
- -> cterm -> thm
- val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
- -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
- -> term -> 'a
-end;
-
-structure Code_Thingol: CODE_THINGOL =
-struct
-
-(** auxiliary **)
-
-fun unfoldl dest x =
- case dest x
- of NONE => (x, [])
- | SOME (x1, x2) =>
- let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end;
-
-fun unfoldr dest x =
- case dest x
- of NONE => ([], x)
- | SOME (x1, x2) =>
- let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end;
-
-
-(** language core - types, terms **)
-
-type vname = string;
-
-datatype dict =
- DictConst of string * dict list list
- | DictVar of string list * (vname * (int * int));
-
-datatype itype =
- `%% of string * itype list
- | ITyVar of vname;
-
-type const = string * ((itype list * dict list list) * itype list (*types of arguments*))
-
-datatype iterm =
- IConst of const
- | IVar of vname
- | `$ of iterm * iterm
- | `|-> of (vname * itype) * iterm
- | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
- (*see also signature*)
-
-val op `$$ = Library.foldl (op `$);
-val op `|--> = Library.foldr (op `|->);
-
-val unfold_app = unfoldl
- (fn op `$ t => SOME t
- | _ => NONE);
-
-val split_abs =
- (fn (v, ty) `|-> (t as ICase (((IVar w, _), [(p, t')]), _)) =>
- if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t)
- | (v, ty) `|-> t => SOME (((v, NONE), ty), t)
- | _ => NONE);
-
-val unfold_abs = unfoldr split_abs;
-
-val split_let =
- (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)
- | _ => NONE);
-
-val unfold_let = unfoldr split_let;
-
-fun unfold_const_app t =
- case unfold_app t
- of (IConst c, ts) => SOME (c, ts)
- | _ => NONE;
-
-fun fold_aiterms f (t as IConst _) = f t
- | fold_aiterms f (t as IVar _) = f t
- | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2
- | fold_aiterms f (t as _ `|-> t') = f t #> fold_aiterms f t'
- | fold_aiterms f (ICase (_, t)) = fold_aiterms f t;
-
-fun fold_constnames f =
- let
- fun add (IConst (c, _)) = f c
- | add _ = I;
- in fold_aiterms add end;
-
-fun fold_varnames f =
- let
- fun add (IVar v) = f v
- | add ((v, _) `|-> _) = f v
- | add _ = I;
- in fold_aiterms add end;
-
-fun fold_unbound_varnames f =
- let
- fun add _ (IConst _) = I
- | add vs (IVar v) = if not (member (op =) vs v) then f v else I
- | add vs (t1 `$ t2) = add vs t1 #> add vs t2
- | add vs ((v, _) `|-> t) = add (insert (op =) v vs) t
- | add vs (ICase (_, t)) = add vs t;
- in add [] end;
-
-fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) =
- let
- fun exists_v t = fold_unbound_varnames (fn w => fn b =>
- b orelse v = w) t false;
- in if v = w andalso forall (fn (t1, t2) =>
- exists_v t1 orelse not (exists_v t2)) ds
- then ((se, ty), ds)
- else ((se, ty), [(IVar v, be)])
- end
- | collapse_let (((v, ty), se), be) =
- ((se, ty), [(IVar v, be)])
-
-fun eta_expand k (c as (_, (_, tys)), ts) =
- let
- val j = length ts;
- val l = k - j;
- val ctxt = (fold o fold_varnames) Name.declare ts Name.context;
- val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys);
- in vs_tys `|--> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end;
-
-fun contains_dictvar t =
- let
- fun contains (DictConst (_, dss)) = (fold o fold) contains dss
- | contains (DictVar _) = K true;
- in
- fold_aiterms
- (fn IConst (_, ((_, dss), _)) => (fold o fold) contains dss | _ => I) t false
- end;
-
-fun locally_monomorphic (IConst _) = false
- | locally_monomorphic (IVar _) = true
- | locally_monomorphic (t `$ _) = locally_monomorphic t
- | locally_monomorphic (_ `|-> t) = locally_monomorphic t
- | locally_monomorphic (ICase ((_, ds), _)) = exists (locally_monomorphic o snd) ds;
-
-
-(** namings **)
-
-(* policies *)
-
-local
- fun thyname_of thy f x = the (AList.lookup (op =) (f x) Markup.theory_nameN);
- fun thyname_of_class thy =
- thyname_of thy (ProofContext.query_class (ProofContext.init thy));
- fun thyname_of_tyco thy =
- thyname_of thy (Type.the_tags (Sign.tsig_of thy));
- fun thyname_of_instance thy inst = case AxClass.arity_property thy inst Markup.theory_nameN
- of [] => error ("no such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
- | thyname :: _ => thyname;
- fun thyname_of_const thy c = case AxClass.class_of_param thy c
- of SOME class => thyname_of_class thy class
- | NONE => (case Code.get_datatype_of_constr thy c
- of SOME dtco => thyname_of_tyco thy dtco
- | NONE => thyname_of thy (Consts.the_tags (Sign.consts_of thy)) c);
- fun purify_base "op &" = "and"
- | purify_base "op |" = "or"
- | purify_base "op -->" = "implies"
- | purify_base "op :" = "member"
- | purify_base "op =" = "eq"
- | purify_base "*" = "product"
- | purify_base "+" = "sum"
- | purify_base s = Name.desymbolize false s;
- fun namify thy get_basename get_thyname name =
- let
- val prefix = get_thyname thy name;
- val base = (purify_base o get_basename) name;
- in Long_Name.append prefix base end;
-in
-
-fun namify_class thy = namify thy Long_Name.base_name thyname_of_class;
-fun namify_classrel thy = namify thy (fn (class1, class2) =>
- Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1) (fn thy => thyname_of_class thy o fst);
- (*order fits nicely with composed projections*)
-fun namify_tyco thy "fun" = "Pure.fun"
- | namify_tyco thy tyco = namify thy Long_Name.base_name thyname_of_tyco tyco;
-fun namify_instance thy = namify thy (fn (class, tyco) =>
- Long_Name.base_name class ^ "_" ^ Long_Name.base_name tyco) thyname_of_instance;
-fun namify_const thy = namify thy Long_Name.base_name thyname_of_const;
-
-end; (* local *)
-
-
-(* data *)
-
-datatype naming = Naming of {
- class: class Symtab.table * Name.context,
- classrel: string Symreltab.table * Name.context,
- tyco: string Symtab.table * Name.context,
- instance: string Symreltab.table * Name.context,
- const: string Symtab.table * Name.context
-}
-
-fun dest_Naming (Naming naming) = naming;
-
-val empty_naming = Naming {
- class = (Symtab.empty, Name.context),
- classrel = (Symreltab.empty, Name.context),
- tyco = (Symtab.empty, Name.context),
- instance = (Symreltab.empty, Name.context),
- const = (Symtab.empty, Name.context)
-};
-
-local
- fun mk_naming (class, classrel, tyco, instance, const) =
- Naming { class = class, classrel = classrel,
- tyco = tyco, instance = instance, const = const };
- fun map_naming f (Naming { class, classrel, tyco, instance, const }) =
- mk_naming (f (class, classrel, tyco, instance, const));
-in
- fun map_class f = map_naming
- (fn (class, classrel, tyco, inst, const) =>
- (f class, classrel, tyco, inst, const));
- fun map_classrel f = map_naming
- (fn (class, classrel, tyco, inst, const) =>
- (class, f classrel, tyco, inst, const));
- fun map_tyco f = map_naming
- (fn (class, classrel, tyco, inst, const) =>
- (class, classrel, f tyco, inst, const));
- fun map_instance f = map_naming
- (fn (class, classrel, tyco, inst, const) =>
- (class, classrel, tyco, f inst, const));
- fun map_const f = map_naming
- (fn (class, classrel, tyco, inst, const) =>
- (class, classrel, tyco, inst, f const));
-end; (*local*)
-
-fun add_variant update (thing, name) (tab, used) =
- let
- val (name', used') = yield_singleton Name.variants name used;
- val tab' = update (thing, name') tab;
- in (tab', used') end;
-
-fun declare thy mapp lookup update namify thing =
- mapp (add_variant update (thing, namify thy thing))
- #> `(fn naming => the (lookup naming thing));
-
-
-(* lookup and declare *)
-
-local
-
-val suffix_class = "class";
-val suffix_classrel = "classrel"
-val suffix_tyco = "tyco";
-val suffix_instance = "inst";
-val suffix_const = "const";
-
-fun add_suffix nsp NONE = NONE
- | add_suffix nsp (SOME name) = SOME (Long_Name.append name nsp);
-
-in
-
-val lookup_class = add_suffix suffix_class
- oo Symtab.lookup o fst o #class o dest_Naming;
-val lookup_classrel = add_suffix suffix_classrel
- oo Symreltab.lookup o fst o #classrel o dest_Naming;
-val lookup_tyco = add_suffix suffix_tyco
- oo Symtab.lookup o fst o #tyco o dest_Naming;
-val lookup_instance = add_suffix suffix_instance
- oo Symreltab.lookup o fst o #instance o dest_Naming;
-val lookup_const = add_suffix suffix_const
- oo Symtab.lookup o fst o #const o dest_Naming;
-
-fun declare_class thy = declare thy map_class
- lookup_class Symtab.update_new namify_class;
-fun declare_classrel thy = declare thy map_classrel
- lookup_classrel Symreltab.update_new namify_classrel;
-fun declare_tyco thy = declare thy map_tyco
- lookup_tyco Symtab.update_new namify_tyco;
-fun declare_instance thy = declare thy map_instance
- lookup_instance Symreltab.update_new namify_instance;
-fun declare_const thy = declare thy map_const
- lookup_const Symtab.update_new namify_const;
-
-fun ensure_declared_const thy const naming =
- case lookup_const naming const
- of SOME const' => (const', naming)
- | NONE => declare_const thy const naming;
-
-val unfold_fun = unfoldr
- (fn "Pure.fun.tyco" `%% [ty1, ty2] => SOME (ty1, ty2)
- | _ => NONE); (*depends on suffix_tyco and namify_tyco!*)
-
-end; (* local *)
-
-
-(** statements, abstract programs **)
-
-type typscheme = (vname * sort) list * itype;
-datatype stmt =
- NoStmt
- | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list)
- | Datatype of string * ((vname * sort) list * (string * itype list) list)
- | Datatypecons of string * string
- | Class of class * (vname * ((class * string) list * (string * itype) list))
- | Classrel of class * class
- | Classparam of string * class
- | Classinst of (class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list);
-
-type program = stmt Graph.T;
-
-fun empty_funs program =
- Graph.fold (fn (name, (Fun (c, (_, [])), _)) => cons c
- | _ => I) program [];
-
-fun map_terms_bottom_up f (t as IConst _) = f t
- | map_terms_bottom_up f (t as IVar _) = f t
- | map_terms_bottom_up f (t1 `$ t2) = f
- (map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2)
- | map_terms_bottom_up f ((v, ty) `|-> t) = f
- ((v, ty) `|-> map_terms_bottom_up f t)
- | map_terms_bottom_up f (ICase (((t, ty), ps), t0)) = f
- (ICase (((map_terms_bottom_up f t, ty), (map o pairself)
- (map_terms_bottom_up f) ps), map_terms_bottom_up f t0));
-
-fun map_terms_stmt f NoStmt = NoStmt
- | map_terms_stmt f (Fun (c, (tysm, eqs))) = Fun (c, (tysm, (map o apfst)
- (fn (ts, t) => (map f ts, f t)) eqs))
- | map_terms_stmt f (stmt as Datatype _) = stmt
- | map_terms_stmt f (stmt as Datatypecons _) = stmt
- | map_terms_stmt f (stmt as Class _) = stmt
- | map_terms_stmt f (stmt as Classrel _) = stmt
- | map_terms_stmt f (stmt as Classparam _) = stmt
- | map_terms_stmt f (Classinst (arity, (superarities, classparms))) =
- Classinst (arity, (superarities, (map o apfst o apsnd) (fn const =>
- case f (IConst const) of IConst const' => const') classparms));
-
-fun is_cons program name = case Graph.get_node program name
- of Datatypecons _ => true
- | _ => false;
-
-fun contr_classparam_typs program name = case Graph.get_node program name
- of Classparam (_, class) => let
- val Class (_, (_, (_, params))) = Graph.get_node program class;
- val SOME ty = AList.lookup (op =) params name;
- val (tys, res_ty) = unfold_fun ty;
- fun no_tyvar (_ `%% tys) = forall no_tyvar tys
- | no_tyvar (ITyVar _) = false;
- in if no_tyvar res_ty
- then map (fn ty => if no_tyvar ty then NONE else SOME ty) tys
- else []
- end
- | _ => [];
-
-
-(** translation kernel **)
-
-(* generic mechanisms *)
-
-fun ensure_stmt lookup declare generate thing (dep, (naming, program)) =
- let
- fun add_dep name = case dep of NONE => I
- | SOME dep => Graph.add_edge (dep, name);
- val (name, naming') = case lookup naming thing
- of SOME name => (name, naming)
- | NONE => declare thing naming;
- in case try (Graph.get_node program) name
- of SOME stmt => program
- |> add_dep name
- |> pair naming'
- |> pair dep
- |> pair name
- | NONE => program
- |> Graph.default_node (name, NoStmt)
- |> add_dep name
- |> pair naming'
- |> curry generate (SOME name)
- ||> snd
- |-> (fn stmt => (apsnd o Graph.map_node name) (K stmt))
- |> pair dep
- |> pair name
- end;
-
-fun not_wellsorted thy thm ty sort e =
- let
- val err_class = Sorts.class_error (Syntax.pp_global thy) e;
- val err_thm = case thm
- of SOME thm => "\n(in code equation " ^ Display.string_of_thm thm ^ ")" | NONE => "";
- val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort "
- ^ Syntax.string_of_sort_global thy sort;
- in error ("Wellsortedness error" ^ err_thm ^ ":\n" ^ err_typ ^ "\n" ^ err_class) end;
-
-
-(* translation *)
-
-fun ensure_tyco thy algbr funcgr tyco =
- let
- val stmt_datatype =
- let
- val (vs, cos) = Code.get_datatype thy tyco;
- in
- fold_map (translate_tyvar_sort thy algbr funcgr) vs
- ##>> fold_map (fn (c, tys) =>
- ensure_const thy algbr funcgr c
- ##>> fold_map (translate_typ thy algbr funcgr) tys) cos
- #>> (fn info => Datatype (tyco, info))
- end;
- in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end
-and ensure_const thy algbr funcgr c =
- let
- fun stmt_datatypecons tyco =
- ensure_tyco thy algbr funcgr tyco
- #>> (fn tyco => Datatypecons (c, tyco));
- fun stmt_classparam class =
- ensure_class thy algbr funcgr class
- #>> (fn class => Classparam (c, class));
- fun stmt_fun ((vs, ty), raw_thms) =
- let
- val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
- then raw_thms
- else (map o apfst) (Code.expand_eta thy 1) raw_thms;
- in
- fold_map (translate_tyvar_sort thy algbr funcgr) vs
- ##>> translate_typ thy algbr funcgr ty
- ##>> fold_map (translate_eq thy algbr funcgr) thms
- #>> (fn info => Fun (c, info))
- end;
- val stmt_const = case Code.get_datatype_of_constr thy c
- of SOME tyco => stmt_datatypecons tyco
- | NONE => (case AxClass.class_of_param thy c
- of SOME class => stmt_classparam class
- | NONE => stmt_fun (Code_Preproc.typ funcgr c, Code_Preproc.eqns funcgr c))
- in ensure_stmt lookup_const (declare_const thy) stmt_const c end
-and ensure_class thy (algbr as (_, algebra)) funcgr class =
- let
- val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
- val cs = #params (AxClass.get_info thy class);
- val stmt_class =
- fold_map (fn superclass => ensure_class thy algbr funcgr superclass
- ##>> ensure_classrel thy algbr funcgr (class, superclass)) superclasses
- ##>> fold_map (fn (c, ty) => ensure_const thy algbr funcgr c
- ##>> translate_typ thy algbr funcgr ty) cs
- #>> (fn info => Class (class, (unprefix "'" Name.aT, info)))
- in ensure_stmt lookup_class (declare_class thy) stmt_class class end
-and ensure_classrel thy algbr funcgr (subclass, superclass) =
- let
- val stmt_classrel =
- ensure_class thy algbr funcgr subclass
- ##>> ensure_class thy algbr funcgr superclass
- #>> Classrel;
- in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (subclass, superclass) end
-and ensure_inst thy (algbr as (_, algebra)) funcgr (class, tyco) =
- let
- val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
- val classparams = these (try (#params o AxClass.get_info thy) class);
- val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]);
- val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class];
- val vs' = map2 (fn (v, sort1) => fn sort2 => (v,
- Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts';
- val arity_typ = Type (tyco, map TFree vs);
- val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs');
- fun translate_superarity superclass =
- ensure_class thy algbr funcgr superclass
- ##>> ensure_classrel thy algbr funcgr (class, superclass)
- ##>> translate_dicts thy algbr funcgr NONE (arity_typ, [superclass])
- #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
- (superclass, (classrel, (inst, dss))));
- fun translate_classparam_inst (c, ty) =
- let
- val c_inst = Const (c, map_type_tfree (K arity_typ') ty);
- val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy c_inst);
- val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd
- o Logic.dest_equals o Thm.prop_of) thm;
- in
- ensure_const thy algbr funcgr c
- ##>> translate_const thy algbr funcgr (SOME thm) c_ty
- #>> (fn (c, IConst c_inst) => ((c, c_inst), (thm, true)))
- end;
- val stmt_inst =
- ensure_class thy algbr funcgr class
- ##>> ensure_tyco thy algbr funcgr tyco
- ##>> fold_map (translate_tyvar_sort thy algbr funcgr) vs
- ##>> fold_map translate_superarity superclasses
- ##>> fold_map translate_classparam_inst classparams
- #>> (fn ((((class, tyco), arity), superarities), classparams) =>
- Classinst ((class, (tyco, arity)), (superarities, classparams)));
- in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end
-and translate_typ thy algbr funcgr (TFree (v, _)) =
- pair (ITyVar (unprefix "'" v))
- | translate_typ thy algbr funcgr (Type (tyco, tys)) =
- ensure_tyco thy algbr funcgr tyco
- ##>> fold_map (translate_typ thy algbr funcgr) tys
- #>> (fn (tyco, tys) => tyco `%% tys)
-and translate_term thy algbr funcgr thm (Const (c, ty)) =
- translate_app thy algbr funcgr thm ((c, ty), [])
- | translate_term thy algbr funcgr thm (Free (v, _)) =
- pair (IVar v)
- | translate_term thy algbr funcgr thm (Abs (abs as (_, ty, _))) =
- let
- val (v, t) = Syntax.variant_abs abs;
- in
- translate_typ thy algbr funcgr ty
- ##>> translate_term thy algbr funcgr thm t
- #>> (fn (ty, t) => (v, ty) `|-> t)
- end
- | translate_term thy algbr funcgr thm (t as _ $ _) =
- case strip_comb t
- of (Const (c, ty), ts) =>
- translate_app thy algbr funcgr thm ((c, ty), ts)
- | (t', ts) =>
- translate_term thy algbr funcgr thm t'
- ##>> fold_map (translate_term thy algbr funcgr thm) ts
- #>> (fn (t, ts) => t `$$ ts)
-and translate_eq thy algbr funcgr (thm, proper) =
- let
- val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals
- o Logic.unvarify o prop_of) thm;
- in
- fold_map (translate_term thy algbr funcgr (SOME thm)) args
- ##>> translate_term thy algbr funcgr (SOME thm) rhs
- #>> rpair (thm, proper)
- end
-and translate_const thy algbr funcgr thm (c, ty) =
- let
- val tys = Sign.const_typargs thy (c, ty);
- val sorts = (map snd o fst o Code_Preproc.typ funcgr) c;
- val tys_args = (fst o Term.strip_type) ty;
- in
- ensure_const thy algbr funcgr c
- ##>> fold_map (translate_typ thy algbr funcgr) tys
- ##>> fold_map (translate_dicts thy algbr funcgr thm) (tys ~~ sorts)
- ##>> fold_map (translate_typ thy algbr funcgr) tys_args
- #>> (fn (((c, tys), iss), tys_args) => IConst (c, ((tys, iss), tys_args)))
- end
-and translate_app_const thy algbr funcgr thm (c_ty, ts) =
- translate_const thy algbr funcgr thm c_ty
- ##>> fold_map (translate_term thy algbr funcgr thm) ts
- #>> (fn (t, ts) => t `$$ ts)
-and translate_case thy algbr funcgr thm (num_args, (t_pos, case_pats)) (c_ty, ts) =
- let
- val (tys, _) = (chop num_args o fst o strip_type o snd) c_ty;
- val t = nth ts t_pos;
- val ty = nth tys t_pos;
- val ts_clause = nth_drop t_pos ts;
- fun mk_clause (co, num_co_args) t =
- let
- val (vs, body) = Term.strip_abs_eta num_co_args t;
- val not_undefined = case body
- of (Const (c, _)) => not (Code.is_undefined thy c)
- | _ => true;
- val pat = list_comb (Const (co, map snd vs ---> ty), map Free vs);
- in (not_undefined, (pat, body)) end;
- val clauses = if null case_pats then let val ([v_ty], body) =
- Term.strip_abs_eta 1 (the_single ts_clause)
- in [(true, (Free v_ty, body))] end
- else map (uncurry mk_clause)
- (AList.make (Code.no_args thy) case_pats ~~ ts_clause);
- fun retermify ty (_, (IVar x, body)) =
- (x, ty) `|-> body
- | retermify _ (_, (pat, body)) =
- let
- val (IConst (_, (_, tys)), ts) = unfold_app pat;
- val vs = map2 (fn IVar x => fn ty => (x, ty)) ts tys;
- in vs `|--> body end;
- fun mk_icase const t ty clauses =
- let
- val (ts1, ts2) = chop t_pos (map (retermify ty) clauses);
- in
- ICase (((t, ty), map_filter (fn (b, d) => if b then SOME d else NONE) clauses),
- const `$$ (ts1 @ t :: ts2))
- end;
- in
- translate_const thy algbr funcgr thm c_ty
- ##>> translate_term thy algbr funcgr thm t
- ##>> translate_typ thy algbr funcgr ty
- ##>> fold_map (fn (b, (pat, body)) => translate_term thy algbr funcgr thm pat
- ##>> translate_term thy algbr funcgr thm body
- #>> pair b) clauses
- #>> (fn (((const, t), ty), ds) => mk_icase const t ty ds)
- end
-and translate_app_case thy algbr funcgr thm (case_scheme as (num_args, _)) ((c, ty), ts) =
- if length ts < num_args then
- let
- val k = length ts;
- val tys = (curry Library.take (num_args - k) o curry Library.drop k o fst o strip_type) ty;
- val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context;
- val vs = Name.names ctxt "a" tys;
- in
- fold_map (translate_typ thy algbr funcgr) tys
- ##>> translate_case thy algbr funcgr thm case_scheme ((c, ty), ts @ map Free vs)
- #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t)
- end
- else if length ts > num_args then
- translate_case thy algbr funcgr thm case_scheme ((c, ty), Library.take (num_args, ts))
- ##>> fold_map (translate_term thy algbr funcgr thm) (Library.drop (num_args, ts))
- #>> (fn (t, ts) => t `$$ ts)
- else
- translate_case thy algbr funcgr thm case_scheme ((c, ty), ts)
-and translate_app thy algbr funcgr thm (c_ty_ts as ((c, _), _)) =
- case Code.get_case_scheme thy c
- of SOME case_scheme => translate_app_case thy algbr funcgr thm case_scheme c_ty_ts
- | NONE => translate_app_const thy algbr funcgr thm c_ty_ts
-and translate_tyvar_sort thy (algbr as (proj_sort, _)) funcgr (v, sort) =
- fold_map (ensure_class thy algbr funcgr) (proj_sort sort)
- #>> (fn sort => (unprefix "'" v, sort))
-and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) =
- let
- val pp = Syntax.pp_global thy;
- datatype typarg =
- Global of (class * string) * typarg list list
- | Local of (class * class) list * (string * (int * sort));
- fun class_relation (Global ((_, tyco), yss), _) class =
- Global ((class, tyco), yss)
- | class_relation (Local (classrels, v), subclass) superclass =
- Local ((subclass, superclass) :: classrels, v);
- fun type_constructor tyco yss class =
- Global ((class, tyco), (map o map) fst yss);
- fun type_variable (TFree (v, sort)) =
- let
- val sort' = proj_sort sort;
- in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
- val typargs = Sorts.of_sort_derivation pp algebra
- {class_relation = class_relation, type_constructor = type_constructor,
- type_variable = type_variable} (ty, proj_sort sort)
- handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e;
- fun mk_dict (Global (inst, yss)) =
- ensure_inst thy algbr funcgr inst
- ##>> (fold_map o fold_map) mk_dict yss
- #>> (fn (inst, dss) => DictConst (inst, dss))
- | mk_dict (Local (classrels, (v, (k, sort)))) =
- fold_map (ensure_classrel thy algbr funcgr) classrels
- #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
- in fold_map mk_dict typargs end;
-
-
-(* store *)
-
-structure Program = CodeDataFun
-(
- type T = naming * program;
- val empty = (empty_naming, Graph.empty);
- fun purge thy cs (naming, program) =
- let
- val names_delete = cs
- |> map_filter (lookup_const naming)
- |> filter (can (Graph.get_node program))
- |> Graph.all_preds program;
- val program' = Graph.del_nodes names_delete program;
- in (naming, program') end;
-);
-
-val cached_program = Program.get;
-
-fun invoke_generation thy (algebra, funcgr) f name =
- Program.change_yield thy (fn naming_program => (NONE, naming_program)
- |> f thy algebra funcgr name
- |-> (fn name => fn (_, naming_program) => (name, naming_program)));
-
-
-(* program generation *)
-
-fun consts_program thy cs =
- let
- fun project_consts cs (naming, program) =
- let
- val cs_all = Graph.all_succs program cs;
- in (cs, (naming, Graph.subgraph (member (op =) cs_all) program)) end;
- fun generate_consts thy algebra funcgr =
- fold_map (ensure_const thy algebra funcgr);
- in
- invoke_generation thy (Code_Preproc.obtain thy cs []) generate_consts cs
- |-> project_consts
- end;
-
-
-(* value evaluation *)
-
-fun ensure_value thy algbr funcgr t =
- let
- val ty = fastype_of t;
- val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =)
- o dest_TFree))) t [];
- val stmt_value =
- fold_map (translate_tyvar_sort thy algbr funcgr) vs
- ##>> translate_typ thy algbr funcgr ty
- ##>> translate_term thy algbr funcgr NONE t
- #>> (fn ((vs, ty), t) => Fun
- (Term.dummy_patternN, ((vs, ty), [(([], t), (Drule.dummy_thm, true))])));
- fun term_value (dep, (naming, program1)) =
- let
- val Fun (_, (vs_ty, [(([], t), _)])) =
- Graph.get_node program1 Term.dummy_patternN;
- val deps = Graph.imm_succs program1 Term.dummy_patternN;
- val program2 = Graph.del_nodes [Term.dummy_patternN] program1;
- val deps_all = Graph.all_succs program2 deps;
- val program3 = Graph.subgraph (member (op =) deps_all) program2;
- in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end;
- in
- ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN
- #> snd
- #> term_value
- end;
-
-fun base_evaluator thy evaluator algebra funcgr vs t =
- let
- val (((naming, program), (((vs', ty'), t'), deps)), _) =
- invoke_generation thy (algebra, funcgr) ensure_value t;
- val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';
- in evaluator naming program ((vs'', (vs', ty')), t') deps end;
-
-fun eval_conv thy prep_sort = Code_Preproc.eval_conv thy prep_sort o base_evaluator thy;
-fun eval thy prep_sort postproc = Code_Preproc.eval thy prep_sort postproc o base_evaluator thy;
-
-
-(** diagnostic commands **)
-
-fun read_const_exprs thy =
- let
- fun consts_of some_thyname =
- let
- val thy' = case some_thyname
- of SOME thyname => ThyInfo.the_theory thyname thy
- | NONE => thy;
- val cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I)
- ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) thy') [];
- fun belongs_here c =
- not (exists (fn thy'' => Sign.declared_const thy'' c) (Theory.parents_of thy'))
- in case some_thyname
- of NONE => cs
- | SOME thyname => filter belongs_here cs
- end;
- fun read_const_expr "*" = ([], consts_of NONE)
- | read_const_expr s = if String.isSuffix ".*" s
- then ([], consts_of (SOME (unsuffix ".*" s)))
- else ([Code.read_const thy s], []);
- in pairself flat o split_list o map read_const_expr end;
-
-fun code_depgr thy consts =
- let
- val (_, eqngr) = Code_Preproc.obtain thy consts [];
- val select = Graph.all_succs eqngr consts;
- in
- eqngr
- |> not (null consts) ? Graph.subgraph (member (op =) select)
- |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy))
- end;
-
-fun code_thms thy = Pretty.writeln o Code_Preproc.pretty thy o code_depgr thy;
-
-fun code_deps thy consts =
- let
- val eqngr = code_depgr thy consts;
- val constss = Graph.strong_conn eqngr;
- val mapping = Symtab.empty |> fold (fn consts => fold (fn const =>
- Symtab.update (const, consts)) consts) constss;
- fun succs consts = consts
- |> maps (Graph.imm_succs eqngr)
- |> subtract (op =) consts
- |> map (the o Symtab.lookup mapping)
- |> distinct (op =);
- val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
- fun namify consts = map (Code.string_of_const thy) consts
- |> commas;
- val prgr = map (fn (consts, constss) =>
- { name = namify consts, ID = namify consts, dir = "", unfold = true,
- path = "", parents = map namify constss }) conn;
- in Present.display_graph prgr end;
-
-local
-
-structure P = OuterParse
-and K = OuterKeyword
-
-fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy;
-fun code_deps_cmd thy = code_deps thy o op @ o read_const_exprs thy;
-
-in
-
-val _ =
- OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag
- (Scan.repeat P.term_group
- >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
- o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
-
-val _ =
- OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag
- (Scan.repeat P.term_group
- >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
- o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
-
-end;
-
-end; (*struct*)
-
-
-structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol;
--- a/src/Tools/nbe.ML Thu Jun 18 17:14:08 2009 +0100
+++ b/src/Tools/nbe.ML Tue Jun 23 15:48:55 2009 +0100
@@ -192,7 +192,7 @@
in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end
and of_iapp match_cont (IConst (c, ((_, dss), _))) ts = constapp c dss ts
| of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound v) ts
- | of_iapp match_cont ((v, _) `|-> t) ts =
+ | of_iapp match_cont ((v, _) `|=> t) ts =
nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound v]) (of_iterm NONE t))) ts
| of_iapp match_cont (ICase (((t, _), cs), t0)) ts =
nbe_apps (ml_cases (of_iterm NONE t)