renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
--- a/src/Doc/ROOT Mon Sep 01 16:17:46 2014 +0200
+++ b/src/Doc/ROOT Mon Sep 01 16:17:46 2014 +0200
@@ -435,4 +435,3 @@
"tutorial.sty"
"typedef.pdf"
"types0.tex"
-
--- a/src/Doc/Tutorial/ToyList/ToyList.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/Doc/Tutorial/ToyList/ToyList.thy Mon Sep 01 16:17:46 2014 +0200
@@ -1,5 +1,5 @@
theory ToyList
-imports Datatype
+imports Old_Datatype
begin
text{*\noindent
--- a/src/Doc/Tutorial/ToyList/ToyList1.txt Mon Sep 01 16:17:46 2014 +0200
+++ b/src/Doc/Tutorial/ToyList/ToyList1.txt Mon Sep 01 16:17:46 2014 +0200
@@ -1,5 +1,5 @@
theory ToyList
-imports Datatype
+imports Old_Datatype
begin
datatype 'a list = Nil ("[]")
--- a/src/Doc/Tutorial/ToyList/ToyList_Test.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/Doc/Tutorial/ToyList/ToyList_Test.thy Mon Sep 01 16:17:46 2014 +0200
@@ -1,5 +1,5 @@
theory ToyList_Test
-imports Datatype
+imports Old_Datatype
begin
ML {*
@@ -10,4 +10,3 @@
*}
end
-
--- a/src/HOL/BNF_Examples/Compat.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/BNF_Examples/Compat.thy Mon Sep 01 16:17:46 2014 +0200
@@ -80,6 +80,6 @@
datatype_new tree = Tree "tree foo"
datatype_compat tree
-ML {* Datatype_Data.get_info @{theory} @{type_name tree} *}
+ML {* Old_Datatype_Data.get_info @{theory} @{type_name tree} *}
end
--- a/src/HOL/BNF_FP_Base.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/BNF_FP_Base.thy Mon Sep 01 16:17:46 2014 +0200
@@ -205,8 +205,8 @@
ML_file "Tools/BNF/bnf_fp_n2m.ML"
ML_file "Tools/BNF/bnf_fp_n2m_sugar.ML"
-ML_file "Tools/Function/size.ML"
-setup Size.setup
+ML_file "Tools/Function/old_size.ML"
+setup Old_Size.setup
lemma size_bool[code]: "size (b\<Colon>bool) = 0"
by (cases b) auto
--- a/src/HOL/Datatype.thy Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,524 +0,0 @@
-(* Title: HOL/Datatype.thy
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
- Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
-*)
-
-header {* Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums *}
-
-theory Datatype
-imports Product_Type Sum_Type Nat
-keywords "datatype" :: thy_decl
-begin
-
-subsection {* The datatype universe *}
-
-definition "Node = {p. EX f x k. p = (f :: nat => 'b + nat, x ::'a + nat) & f k = Inr 0}"
-
-typedef ('a, 'b) node = "Node :: ((nat => 'b + nat) * ('a + nat)) set"
- morphisms Rep_Node Abs_Node
- unfolding Node_def by auto
-
-text{*Datatypes will be represented by sets of type @{text node}*}
-
-type_synonym 'a item = "('a, unit) node set"
-type_synonym ('a, 'b) dtree = "('a, 'b) node set"
-
-consts
- Push :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))"
-
- Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node"
- ndepth :: "('a, 'b) node => nat"
-
- Atom :: "('a + nat) => ('a, 'b) dtree"
- Leaf :: "'a => ('a, 'b) dtree"
- Numb :: "nat => ('a, 'b) dtree"
- Scons :: "[('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree"
- In0 :: "('a, 'b) dtree => ('a, 'b) dtree"
- In1 :: "('a, 'b) dtree => ('a, 'b) dtree"
- Lim :: "('b => ('a, 'b) dtree) => ('a, 'b) dtree"
-
- ntrunc :: "[nat, ('a, 'b) dtree] => ('a, 'b) dtree"
-
- uprod :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set"
- usum :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set"
-
- Split :: "[[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
- Case :: "[[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
-
- dprod :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set]
- => (('a, 'b) dtree * ('a, 'b) dtree)set"
- dsum :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set]
- => (('a, 'b) dtree * ('a, 'b) dtree)set"
-
-
-defs
-
- Push_Node_def: "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))"
-
- (*crude "lists" of nats -- needed for the constructions*)
- Push_def: "Push == (%b h. case_nat b h)"
-
- (** operations on S-expressions -- sets of nodes **)
-
- (*S-expression constructors*)
- Atom_def: "Atom == (%x. {Abs_Node((%k. Inr 0, x))})"
- Scons_def: "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)"
-
- (*Leaf nodes, with arbitrary or nat labels*)
- Leaf_def: "Leaf == Atom o Inl"
- Numb_def: "Numb == Atom o Inr"
-
- (*Injections of the "disjoint sum"*)
- In0_def: "In0(M) == Scons (Numb 0) M"
- In1_def: "In1(M) == Scons (Numb 1) M"
-
- (*Function spaces*)
- Lim_def: "Lim f == Union {z. ? x. z = Push_Node (Inl x) ` (f x)}"
-
- (*the set of nodes with depth less than k*)
- ndepth_def: "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)"
- ntrunc_def: "ntrunc k N == {n. n:N & ndepth(n)<k}"
-
- (*products and sums for the "universe"*)
- uprod_def: "uprod A B == UN x:A. UN y:B. { Scons x y }"
- usum_def: "usum A B == In0`A Un In1`B"
-
- (*the corresponding eliminators*)
- Split_def: "Split c M == THE u. EX x y. M = Scons x y & u = c x y"
-
- Case_def: "Case c d M == THE u. (EX x . M = In0(x) & u = c(x))
- | (EX y . M = In1(y) & u = d(y))"
-
-
- (** equality for the "universe" **)
-
- dprod_def: "dprod r s == UN (x,x'):r. UN (y,y'):s. {(Scons x y, Scons x' y')}"
-
- dsum_def: "dsum r s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un
- (UN (y,y'):s. {(In1(y),In1(y'))})"
-
-
-
-lemma apfst_convE:
- "[| q = apfst f p; !!x y. [| p = (x,y); q = (f(x),y) |] ==> R
- |] ==> R"
-by (force simp add: apfst_def)
-
-(** Push -- an injection, analogous to Cons on lists **)
-
-lemma Push_inject1: "Push i f = Push j g ==> i=j"
-apply (simp add: Push_def fun_eq_iff)
-apply (drule_tac x=0 in spec, simp)
-done
-
-lemma Push_inject2: "Push i f = Push j g ==> f=g"
-apply (auto simp add: Push_def fun_eq_iff)
-apply (drule_tac x="Suc x" in spec, simp)
-done
-
-lemma Push_inject:
- "[| Push i f =Push j g; [| i=j; f=g |] ==> P |] ==> P"
-by (blast dest: Push_inject1 Push_inject2)
-
-lemma Push_neq_K0: "Push (Inr (Suc k)) f = (%z. Inr 0) ==> P"
-by (auto simp add: Push_def fun_eq_iff split: nat.split_asm)
-
-lemmas Abs_Node_inj = Abs_Node_inject [THEN [2] rev_iffD1]
-
-
-(*** Introduction rules for Node ***)
-
-lemma Node_K0_I: "(%k. Inr 0, a) : Node"
-by (simp add: Node_def)
-
-lemma Node_Push_I: "p: Node ==> apfst (Push i) p : Node"
-apply (simp add: Node_def Push_def)
-apply (fast intro!: apfst_conv nat.case(2)[THEN trans])
-done
-
-
-subsection{*Freeness: Distinctness of Constructors*}
-
-(** Scons vs Atom **)
-
-lemma Scons_not_Atom [iff]: "Scons M N \<noteq> Atom(a)"
-unfolding Atom_def Scons_def Push_Node_def One_nat_def
-by (blast intro: Node_K0_I Rep_Node [THEN Node_Push_I]
- dest!: Abs_Node_inj
- elim!: apfst_convE sym [THEN Push_neq_K0])
-
-lemmas Atom_not_Scons [iff] = Scons_not_Atom [THEN not_sym]
-
-
-(*** Injectiveness ***)
-
-(** Atomic nodes **)
-
-lemma inj_Atom: "inj(Atom)"
-apply (simp add: Atom_def)
-apply (blast intro!: inj_onI Node_K0_I dest!: Abs_Node_inj)
-done
-lemmas Atom_inject = inj_Atom [THEN injD]
-
-lemma Atom_Atom_eq [iff]: "(Atom(a)=Atom(b)) = (a=b)"
-by (blast dest!: Atom_inject)
-
-lemma inj_Leaf: "inj(Leaf)"
-apply (simp add: Leaf_def o_def)
-apply (rule inj_onI)
-apply (erule Atom_inject [THEN Inl_inject])
-done
-
-lemmas Leaf_inject [dest!] = inj_Leaf [THEN injD]
-
-lemma inj_Numb: "inj(Numb)"
-apply (simp add: Numb_def o_def)
-apply (rule inj_onI)
-apply (erule Atom_inject [THEN Inr_inject])
-done
-
-lemmas Numb_inject [dest!] = inj_Numb [THEN injD]
-
-
-(** Injectiveness of Push_Node **)
-
-lemma Push_Node_inject:
- "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P
- |] ==> P"
-apply (simp add: Push_Node_def)
-apply (erule Abs_Node_inj [THEN apfst_convE])
-apply (rule Rep_Node [THEN Node_Push_I])+
-apply (erule sym [THEN apfst_convE])
-apply (blast intro: Rep_Node_inject [THEN iffD1] trans sym elim!: Push_inject)
-done
-
-
-(** Injectiveness of Scons **)
-
-lemma Scons_inject_lemma1: "Scons M N <= Scons M' N' ==> M<=M'"
-unfolding Scons_def One_nat_def
-by (blast dest!: Push_Node_inject)
-
-lemma Scons_inject_lemma2: "Scons M N <= Scons M' N' ==> N<=N'"
-unfolding Scons_def One_nat_def
-by (blast dest!: Push_Node_inject)
-
-lemma Scons_inject1: "Scons M N = Scons M' N' ==> M=M'"
-apply (erule equalityE)
-apply (iprover intro: equalityI Scons_inject_lemma1)
-done
-
-lemma Scons_inject2: "Scons M N = Scons M' N' ==> N=N'"
-apply (erule equalityE)
-apply (iprover intro: equalityI Scons_inject_lemma2)
-done
-
-lemma Scons_inject:
- "[| Scons M N = Scons M' N'; [| M=M'; N=N' |] ==> P |] ==> P"
-by (iprover dest: Scons_inject1 Scons_inject2)
-
-lemma Scons_Scons_eq [iff]: "(Scons M N = Scons M' N') = (M=M' & N=N')"
-by (blast elim!: Scons_inject)
-
-(*** Distinctness involving Leaf and Numb ***)
-
-(** Scons vs Leaf **)
-
-lemma Scons_not_Leaf [iff]: "Scons M N \<noteq> Leaf(a)"
-unfolding Leaf_def o_def by (rule Scons_not_Atom)
-
-lemmas Leaf_not_Scons [iff] = Scons_not_Leaf [THEN not_sym]
-
-(** Scons vs Numb **)
-
-lemma Scons_not_Numb [iff]: "Scons M N \<noteq> Numb(k)"
-unfolding Numb_def o_def by (rule Scons_not_Atom)
-
-lemmas Numb_not_Scons [iff] = Scons_not_Numb [THEN not_sym]
-
-
-(** Leaf vs Numb **)
-
-lemma Leaf_not_Numb [iff]: "Leaf(a) \<noteq> Numb(k)"
-by (simp add: Leaf_def Numb_def)
-
-lemmas Numb_not_Leaf [iff] = Leaf_not_Numb [THEN not_sym]
-
-
-(*** ndepth -- the depth of a node ***)
-
-lemma ndepth_K0: "ndepth (Abs_Node(%k. Inr 0, x)) = 0"
-by (simp add: ndepth_def Node_K0_I [THEN Abs_Node_inverse] Least_equality)
-
-lemma ndepth_Push_Node_aux:
- "case_nat (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"
-apply (induct_tac "k", auto)
-apply (erule Least_le)
-done
-
-lemma ndepth_Push_Node:
- "ndepth (Push_Node (Inr (Suc i)) n) = Suc(ndepth(n))"
-apply (insert Rep_Node [of n, unfolded Node_def])
-apply (auto simp add: ndepth_def Push_Node_def
- Rep_Node [THEN Node_Push_I, THEN Abs_Node_inverse])
-apply (rule Least_equality)
-apply (auto simp add: Push_def ndepth_Push_Node_aux)
-apply (erule LeastI)
-done
-
-
-(*** ntrunc applied to the various node sets ***)
-
-lemma ntrunc_0 [simp]: "ntrunc 0 M = {}"
-by (simp add: ntrunc_def)
-
-lemma ntrunc_Atom [simp]: "ntrunc (Suc k) (Atom a) = Atom(a)"
-by (auto simp add: Atom_def ntrunc_def ndepth_K0)
-
-lemma ntrunc_Leaf [simp]: "ntrunc (Suc k) (Leaf a) = Leaf(a)"
-unfolding Leaf_def o_def by (rule ntrunc_Atom)
-
-lemma ntrunc_Numb [simp]: "ntrunc (Suc k) (Numb i) = Numb(i)"
-unfolding Numb_def o_def by (rule ntrunc_Atom)
-
-lemma ntrunc_Scons [simp]:
- "ntrunc (Suc k) (Scons M N) = Scons (ntrunc k M) (ntrunc k N)"
-unfolding Scons_def ntrunc_def One_nat_def
-by (auto simp add: ndepth_Push_Node)
-
-
-
-(** Injection nodes **)
-
-lemma ntrunc_one_In0 [simp]: "ntrunc (Suc 0) (In0 M) = {}"
-apply (simp add: In0_def)
-apply (simp add: Scons_def)
-done
-
-lemma ntrunc_In0 [simp]: "ntrunc (Suc(Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"
-by (simp add: In0_def)
-
-lemma ntrunc_one_In1 [simp]: "ntrunc (Suc 0) (In1 M) = {}"
-apply (simp add: In1_def)
-apply (simp add: Scons_def)
-done
-
-lemma ntrunc_In1 [simp]: "ntrunc (Suc(Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"
-by (simp add: In1_def)
-
-
-subsection{*Set Constructions*}
-
-
-(*** Cartesian Product ***)
-
-lemma uprodI [intro!]: "[| M:A; N:B |] ==> Scons M N : uprod A B"
-by (simp add: uprod_def)
-
-(*The general elimination rule*)
-lemma uprodE [elim!]:
- "[| c : uprod A B;
- !!x y. [| x:A; y:B; c = Scons x y |] ==> P
- |] ==> P"
-by (auto simp add: uprod_def)
-
-
-(*Elimination of a pair -- introduces no eigenvariables*)
-lemma uprodE2: "[| Scons M N : uprod A B; [| M:A; N:B |] ==> P |] ==> P"
-by (auto simp add: uprod_def)
-
-
-(*** Disjoint Sum ***)
-
-lemma usum_In0I [intro]: "M:A ==> In0(M) : usum A B"
-by (simp add: usum_def)
-
-lemma usum_In1I [intro]: "N:B ==> In1(N) : usum A B"
-by (simp add: usum_def)
-
-lemma usumE [elim!]:
- "[| u : usum A B;
- !!x. [| x:A; u=In0(x) |] ==> P;
- !!y. [| y:B; u=In1(y) |] ==> P
- |] ==> P"
-by (auto simp add: usum_def)
-
-
-(** Injection **)
-
-lemma In0_not_In1 [iff]: "In0(M) \<noteq> In1(N)"
-unfolding In0_def In1_def One_nat_def by auto
-
-lemmas In1_not_In0 [iff] = In0_not_In1 [THEN not_sym]
-
-lemma In0_inject: "In0(M) = In0(N) ==> M=N"
-by (simp add: In0_def)
-
-lemma In1_inject: "In1(M) = In1(N) ==> M=N"
-by (simp add: In1_def)
-
-lemma In0_eq [iff]: "(In0 M = In0 N) = (M=N)"
-by (blast dest!: In0_inject)
-
-lemma In1_eq [iff]: "(In1 M = In1 N) = (M=N)"
-by (blast dest!: In1_inject)
-
-lemma inj_In0: "inj In0"
-by (blast intro!: inj_onI)
-
-lemma inj_In1: "inj In1"
-by (blast intro!: inj_onI)
-
-
-(*** Function spaces ***)
-
-lemma Lim_inject: "Lim f = Lim g ==> f = g"
-apply (simp add: Lim_def)
-apply (rule ext)
-apply (blast elim!: Push_Node_inject)
-done
-
-
-(*** proving equality of sets and functions using ntrunc ***)
-
-lemma ntrunc_subsetI: "ntrunc k M <= M"
-by (auto simp add: ntrunc_def)
-
-lemma ntrunc_subsetD: "(!!k. ntrunc k M <= N) ==> M<=N"
-by (auto simp add: ntrunc_def)
-
-(*A generalized form of the take-lemma*)
-lemma ntrunc_equality: "(!!k. ntrunc k M = ntrunc k N) ==> M=N"
-apply (rule equalityI)
-apply (rule_tac [!] ntrunc_subsetD)
-apply (rule_tac [!] ntrunc_subsetI [THEN [2] subset_trans], auto)
-done
-
-lemma ntrunc_o_equality:
- "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"
-apply (rule ntrunc_equality [THEN ext])
-apply (simp add: fun_eq_iff)
-done
-
-
-(*** Monotonicity ***)
-
-lemma uprod_mono: "[| A<=A'; B<=B' |] ==> uprod A B <= uprod A' B'"
-by (simp add: uprod_def, blast)
-
-lemma usum_mono: "[| A<=A'; B<=B' |] ==> usum A B <= usum A' B'"
-by (simp add: usum_def, blast)
-
-lemma Scons_mono: "[| M<=M'; N<=N' |] ==> Scons M N <= Scons M' N'"
-by (simp add: Scons_def, blast)
-
-lemma In0_mono: "M<=N ==> In0(M) <= In0(N)"
-by (simp add: In0_def Scons_mono)
-
-lemma In1_mono: "M<=N ==> In1(M) <= In1(N)"
-by (simp add: In1_def Scons_mono)
-
-
-(*** Split and Case ***)
-
-lemma Split [simp]: "Split c (Scons M N) = c M N"
-by (simp add: Split_def)
-
-lemma Case_In0 [simp]: "Case c d (In0 M) = c(M)"
-by (simp add: Case_def)
-
-lemma Case_In1 [simp]: "Case c d (In1 N) = d(N)"
-by (simp add: Case_def)
-
-
-
-(**** UN x. B(x) rules ****)
-
-lemma ntrunc_UN1: "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))"
-by (simp add: ntrunc_def, blast)
-
-lemma Scons_UN1_x: "Scons (UN x. f x) M = (UN x. Scons (f x) M)"
-by (simp add: Scons_def, blast)
-
-lemma Scons_UN1_y: "Scons M (UN x. f x) = (UN x. Scons M (f x))"
-by (simp add: Scons_def, blast)
-
-lemma In0_UN1: "In0(UN x. f(x)) = (UN x. In0(f(x)))"
-by (simp add: In0_def Scons_UN1_y)
-
-lemma In1_UN1: "In1(UN x. f(x)) = (UN x. In1(f(x)))"
-by (simp add: In1_def Scons_UN1_y)
-
-
-(*** Equality for Cartesian Product ***)
-
-lemma dprodI [intro!]:
- "[| (M,M'):r; (N,N'):s |] ==> (Scons M N, Scons M' N') : dprod r s"
-by (auto simp add: dprod_def)
-
-(*The general elimination rule*)
-lemma dprodE [elim!]:
- "[| c : dprod r s;
- !!x y x' y'. [| (x,x') : r; (y,y') : s;
- c = (Scons x y, Scons x' y') |] ==> P
- |] ==> P"
-by (auto simp add: dprod_def)
-
-
-(*** Equality for Disjoint Sum ***)
-
-lemma dsum_In0I [intro]: "(M,M'):r ==> (In0(M), In0(M')) : dsum r s"
-by (auto simp add: dsum_def)
-
-lemma dsum_In1I [intro]: "(N,N'):s ==> (In1(N), In1(N')) : dsum r s"
-by (auto simp add: dsum_def)
-
-lemma dsumE [elim!]:
- "[| w : dsum r s;
- !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P;
- !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P
- |] ==> P"
-by (auto simp add: dsum_def)
-
-
-(*** Monotonicity ***)
-
-lemma dprod_mono: "[| r<=r'; s<=s' |] ==> dprod r s <= dprod r' s'"
-by blast
-
-lemma dsum_mono: "[| r<=r'; s<=s' |] ==> dsum r s <= dsum r' s'"
-by blast
-
-
-(*** Bounding theorems ***)
-
-lemma dprod_Sigma: "(dprod (A <*> B) (C <*> D)) <= (uprod A C) <*> (uprod B D)"
-by blast
-
-lemmas dprod_subset_Sigma = subset_trans [OF dprod_mono dprod_Sigma]
-
-(*Dependent version*)
-lemma dprod_subset_Sigma2:
- "(dprod (Sigma A B) (Sigma C D)) <=
- Sigma (uprod A C) (Split (%x y. uprod (B x) (D y)))"
-by auto
-
-lemma dsum_Sigma: "(dsum (A <*> B) (C <*> D)) <= (usum A C) <*> (usum B D)"
-by blast
-
-lemmas dsum_subset_Sigma = subset_trans [OF dsum_mono dsum_Sigma]
-
-
-text {* hides popular names *}
-hide_type (open) node item
-hide_const (open) Push Node Atom Leaf Numb Lim Split Case
-
-ML_file "Tools/Datatype/datatype.ML"
-
-ML_file "Tools/inductive_realizer.ML"
-setup InductiveRealizer.setup
-
-ML_file "Tools/Datatype/datatype_realizer.ML"
-setup Datatype_Realizer.setup
-
-end
--- a/src/HOL/Extraction.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Extraction.thy Mon Sep 01 16:17:46 2014 +0200
@@ -5,7 +5,7 @@
header {* Program extraction for HOL *}
theory Extraction
-imports Datatype Option
+imports Option
begin
ML_file "Tools/rewrite_hol_proof.ML"
--- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Mon Sep 01 16:17:46 2014 +0200
@@ -113,7 +113,7 @@
: (term list * term list) =
let
val Ts = map snd args
- val ns = Name.variant_list taken (Datatype_Prop.make_tnames Ts)
+ val ns = Name.variant_list taken (Old_Datatype_Prop.make_tnames Ts)
val vs = map Free (ns ~~ Ts)
val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs))
in
@@ -167,7 +167,7 @@
fun vars_of args =
let
val Ts = map snd args
- val ns = Datatype_Prop.make_tnames Ts
+ val ns = Old_Datatype_Prop.make_tnames Ts
in
map Free (ns ~~ Ts)
end
@@ -409,7 +409,7 @@
val tns = map fst (Term.add_tfreesT lhsT [])
val resultT = TFree (singleton (Name.variant_list tns) "'t", @{sort pcpo})
fun fTs T = map (fn (_, args) => map snd args -->> T) spec
- val fns = Datatype_Prop.indexify_names (map (K "f") spec)
+ val fns = Old_Datatype_Prop.indexify_names (map (K "f") spec)
val fs = map Free (fns ~~ fTs resultT)
fun caseT T = fTs T -->> (lhsT ->> T)
@@ -424,7 +424,7 @@
fun one_con f (_, args) =
let
val Ts = map snd args
- val ns = Name.variant_list fns (Datatype_Prop.make_tnames Ts)
+ val ns = Name.variant_list fns (Old_Datatype_Prop.make_tnames Ts)
val vs = map Free (ns ~~ Ts)
in
lambda_args (map fst args ~~ vs) (list_ccomb (f, vs))
@@ -606,7 +606,7 @@
fun sel_apps_of (i, (con, args: (bool * term option * typ) list)) =
let
val Ts : typ list = map #3 args
- val ns : string list = Datatype_Prop.make_tnames Ts
+ val ns : string list = Old_Datatype_Prop.make_tnames Ts
val vs : term list = map Free (ns ~~ Ts)
val con_app : term = list_ccomb (con, vs)
val vs' : (bool * term) list = map #1 args ~~ vs
--- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Mon Sep 01 16:17:46 2014 +0200
@@ -63,7 +63,7 @@
fun prove_take_app (con_const, args) =
let
val Ts = map snd args
- val ns = Name.variant_list ["n"] (Datatype_Prop.make_tnames Ts)
+ val ns = Name.variant_list ["n"] (Old_Datatype_Prop.make_tnames Ts)
val vs = map Free (ns ~~ Ts)
val lhs = mk_capply (take_const $ n', list_ccomb (con_const, vs))
val rhs = list_ccomb (con_const, map2 (map_of_arg thy) vs Ts)
@@ -108,8 +108,8 @@
val {take_consts, take_induct_thms, ...} = take_info
val newTs = map #absT iso_infos
- val P_names = Datatype_Prop.indexify_names (map (K "P") newTs)
- val x_names = Datatype_Prop.indexify_names (map (K "x") newTs)
+ val P_names = Old_Datatype_Prop.indexify_names (map (K "P") newTs)
+ val x_names = Old_Datatype_Prop.indexify_names (map (K "x") newTs)
val P_types = map (fn T => T --> HOLogic.boolT) newTs
val Ps = map Free (P_names ~~ P_types)
val xs = map Free (x_names ~~ newTs)
@@ -118,7 +118,7 @@
fun con_assm defined p (con, args) =
let
val Ts = map snd args
- val ns = Name.variant_list P_names (Datatype_Prop.make_tnames Ts)
+ val ns = Name.variant_list P_names (Old_Datatype_Prop.make_tnames Ts)
val vs = map Free (ns ~~ Ts)
val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs))
fun ind_hyp (v, T) t =
@@ -255,7 +255,7 @@
val {take_consts, take_0_thms, take_lemma_thms, ...} = take_info
- val R_names = Datatype_Prop.indexify_names (map (K "R") newTs)
+ val R_names = Old_Datatype_Prop.indexify_names (map (K "R") newTs)
val R_types = map (fn T => T --> T --> boolT) newTs
val Rs = map Free (R_names ~~ R_types)
val n = Free ("n", natT)
@@ -272,7 +272,7 @@
fun one_con T (con, args) =
let
val Ts = map snd args
- val ns1 = Name.variant_list reserved (Datatype_Prop.make_tnames Ts)
+ val ns1 = Name.variant_list reserved (Old_Datatype_Prop.make_tnames Ts)
val ns2 = map (fn n => n^"'") ns1
val vs1 = map Free (ns1 ~~ Ts)
val vs2 = map Free (ns2 ~~ Ts)
--- a/src/HOL/HOLCF/ex/Pattern_Match.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/HOLCF/ex/Pattern_Match.thy Mon Sep 01 16:17:46 2014 +0200
@@ -423,7 +423,7 @@
: (term list * term list) =
let
val Ts = map snd args;
- val ns = Name.variant_list taken (Datatype_Prop.make_tnames Ts);
+ val ns = Name.variant_list taken (Old_Datatype_Prop.make_tnames Ts);
val vs = map Free (ns ~~ Ts);
val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs));
in
@@ -473,10 +473,10 @@
val Ts = map snd args;
val Vs =
(map (K "'t") args)
- |> Datatype_Prop.indexify_names
+ |> Old_Datatype_Prop.indexify_names
|> Name.variant_list tns
|> map (fn t => TFree (t, @{sort pcpo}));
- val patNs = Datatype_Prop.indexify_names (map (K "pat") args);
+ val patNs = Old_Datatype_Prop.indexify_names (map (K "pat") args);
val patTs = map2 (fn T => fn V => T ->> mk_matchT V) Ts Vs;
val pats = map Free (patNs ~~ patTs);
val fail = mk_fail (mk_tupleT Vs);
@@ -539,10 +539,10 @@
val Ts = map snd args;
val Vs =
(map (K "'t") args)
- |> Datatype_Prop.indexify_names
+ |> Old_Datatype_Prop.indexify_names
|> Name.variant_list (rn::tns)
|> map (fn t => TFree (t, @{sort pcpo}));
- val patNs = Datatype_Prop.indexify_names (map (K "pat") args);
+ val patNs = Old_Datatype_Prop.indexify_names (map (K "pat") args);
val patTs = map2 (fn T => fn V => T ->> mk_matchT V) Ts Vs;
val pats = map Free (patNs ~~ patTs);
val k = Free ("rhs", mk_tupleT Vs ->> R);
--- a/src/HOL/Induct/SList.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Induct/SList.thy Mon Sep 01 16:17:46 2014 +0200
@@ -59,8 +59,8 @@
morphisms Rep_List Abs_List
unfolding List_def by (blast intro: list.NIL_I)
-abbreviation "Case == Datatype.Case"
-abbreviation "Split == Datatype.Split"
+abbreviation "Case == Old_Datatype.Case"
+abbreviation "Split == Old_Datatype.Split"
definition
List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b" where
--- a/src/HOL/Induct/Sexp.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Induct/Sexp.thy Mon Sep 01 16:17:46 2014 +0200
@@ -10,9 +10,9 @@
imports Main
begin
-type_synonym 'a item = "'a Datatype.item"
-abbreviation "Leaf == Datatype.Leaf"
-abbreviation "Numb == Datatype.Numb"
+type_synonym 'a item = "'a Old_Datatype.item"
+abbreviation "Leaf == Old_Datatype.Leaf"
+abbreviation "Numb == Old_Datatype.Numb"
inductive_set
sexp :: "'a item set"
--- a/src/HOL/Inductive.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Inductive.thy Mon Sep 01 16:17:46 2014 +0200
@@ -271,12 +271,12 @@
text {* Package setup. *}
-ML_file "Tools/Datatype/datatype_aux.ML"
-ML_file "Tools/Datatype/datatype_prop.ML"
-ML_file "Tools/Datatype/datatype_data.ML" setup Datatype_Data.setup
-ML_file "Tools/Datatype/rep_datatype.ML"
-ML_file "Tools/Datatype/datatype_codegen.ML"
-ML_file "Tools/Datatype/primrec.ML"
+ML_file "Tools/Old_Datatype/old_datatype_aux.ML"
+ML_file "Tools/Old_Datatype/old_datatype_prop.ML"
+ML_file "Tools/Old_Datatype/old_datatype_data.ML" setup Old_Datatype_Data.setup
+ML_file "Tools/Old_Datatype/old_rep_datatype.ML"
+ML_file "Tools/Old_Datatype/old_datatype_codegen.ML"
+ML_file "Tools/Old_Datatype/old_primrec.ML"
ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML"
ML_file "Tools/BNF/bnf_lfp_rec_sugar.ML"
--- a/src/HOL/Library/Countable.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Library/Countable.thy Mon Sep 01 16:17:46 2014 +0200
@@ -173,29 +173,29 @@
subsection {* Automatically proving countability of datatypes *}
-inductive finite_item :: "'a Datatype.item \<Rightarrow> bool" where
+inductive finite_item :: "'a Old_Datatype.item \<Rightarrow> bool" where
undefined: "finite_item undefined"
-| In0: "finite_item x \<Longrightarrow> finite_item (Datatype.In0 x)"
-| In1: "finite_item x \<Longrightarrow> finite_item (Datatype.In1 x)"
-| Leaf: "finite_item (Datatype.Leaf a)"
-| Scons: "\<lbrakk>finite_item x; finite_item y\<rbrakk> \<Longrightarrow> finite_item (Datatype.Scons x y)"
+| In0: "finite_item x \<Longrightarrow> finite_item (Old_Datatype.In0 x)"
+| In1: "finite_item x \<Longrightarrow> finite_item (Old_Datatype.In1 x)"
+| Leaf: "finite_item (Old_Datatype.Leaf a)"
+| Scons: "\<lbrakk>finite_item x; finite_item y\<rbrakk> \<Longrightarrow> finite_item (Old_Datatype.Scons x y)"
function
- nth_item :: "nat \<Rightarrow> ('a::countable) Datatype.item"
+ nth_item :: "nat \<Rightarrow> ('a::countable) Old_Datatype.item"
where
"nth_item 0 = undefined"
| "nth_item (Suc n) =
(case sum_decode n of
Inl i \<Rightarrow>
(case sum_decode i of
- Inl j \<Rightarrow> Datatype.In0 (nth_item j)
- | Inr j \<Rightarrow> Datatype.In1 (nth_item j))
+ Inl j \<Rightarrow> Old_Datatype.In0 (nth_item j)
+ | Inr j \<Rightarrow> Old_Datatype.In1 (nth_item j))
| Inr i \<Rightarrow>
(case sum_decode i of
- Inl j \<Rightarrow> Datatype.Leaf (from_nat j)
+ Inl j \<Rightarrow> Old_Datatype.Leaf (from_nat j)
| Inr j \<Rightarrow>
(case prod_decode j of
- (a, b) \<Rightarrow> Datatype.Scons (nth_item a) (nth_item b))))"
+ (a, b) \<Rightarrow> Old_Datatype.Scons (nth_item a) (nth_item b))))"
by pat_completeness auto
lemma le_sum_encode_Inl: "x \<le> y \<Longrightarrow> x \<le> sum_encode (Inl y)"
@@ -218,33 +218,31 @@
next
case (In0 x)
then obtain n where "nth_item n = x" by fast
- hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inl n)))))
- = Datatype.In0 x" by simp
+ hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inl n))))) = Old_Datatype.In0 x" by simp
thus ?case ..
next
case (In1 x)
then obtain n where "nth_item n = x" by fast
- hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inr n)))))
- = Datatype.In1 x" by simp
+ hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inr n))))) = Old_Datatype.In1 x" by simp
thus ?case ..
next
case (Leaf a)
- have "nth_item (Suc (sum_encode (Inr (sum_encode (Inl (to_nat a))))))
- = Datatype.Leaf a" by simp
+ have "nth_item (Suc (sum_encode (Inr (sum_encode (Inl (to_nat a)))))) = Old_Datatype.Leaf a"
+ by simp
thus ?case ..
next
case (Scons x y)
then obtain i j where "nth_item i = x" and "nth_item j = y" by fast
hence "nth_item
- (Suc (sum_encode (Inr (sum_encode (Inr (prod_encode (i, j)))))))
- = Datatype.Scons x y" by simp
+ (Suc (sum_encode (Inr (sum_encode (Inr (prod_encode (i, j))))))) = Old_Datatype.Scons x y"
+ by simp
thus ?case ..
qed
theorem countable_datatype:
- fixes Rep :: "'b \<Rightarrow> ('a::countable) Datatype.item"
- fixes Abs :: "('a::countable) Datatype.item \<Rightarrow> 'b"
- fixes rep_set :: "('a::countable) Datatype.item \<Rightarrow> bool"
+ fixes Rep :: "'b \<Rightarrow> ('a::countable) Old_Datatype.item"
+ fixes Abs :: "('a::countable) Old_Datatype.item \<Rightarrow> 'b"
+ fixes rep_set :: "('a::countable) Old_Datatype.item \<Rightarrow> bool"
assumes type: "type_definition Rep Abs (Collect rep_set)"
assumes finite_item: "\<And>x. rep_set x \<Longrightarrow> finite_item x"
shows "OFCLASS('b, countable_class)"
--- a/src/HOL/Library/Old_SMT/old_smt_normalize.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Library/Old_SMT/old_smt_normalize.ML Mon Sep 01 16:17:46 2014 +0200
@@ -158,7 +158,7 @@
| @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
| _ => raise CTERM ("no equation", [ct]))
- fun get_constrs thy (Type (n, _)) = these (Datatype_Data.get_constrs thy n)
+ fun get_constrs thy (Type (n, _)) = these (Old_Datatype_Data.get_constrs thy n)
| get_constrs _ _ = []
fun is_constr thy (n, T) =
--- a/src/HOL/Nominal/nominal_atoms.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Nominal/nominal_atoms.ML Mon Sep 01 16:17:46 2014 +0200
@@ -100,9 +100,10 @@
val (_,thy1) =
fold_map (fn ak => fn thy =>
let val dt = ((Binding.name ak, [], NoSyn), [(Binding.name ak, [@{typ nat}], NoSyn)])
- val (dt_names, thy1) = Datatype.add_datatype Datatype_Aux.default_config [dt] thy;
+ val (dt_names, thy1) =
+ Old_Datatype.add_datatype Old_Datatype_Aux.default_config [dt] thy;
- val injects = maps (#inject o Datatype_Data.the_info thy1) dt_names;
+ val injects = maps (#inject o Old_Datatype_Data.the_info thy1) dt_names;
val ak_type = Type (Sign.intern_type thy1 ak,[])
val ak_sign = Sign.intern_const thy1 ak
@@ -190,7 +191,7 @@
val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c)
in
thy' |>
- Primrec.add_primrec_global
+ Old_Primrec.add_primrec_global
[(Binding.name swap_name, SOME swapT, NoSyn)]
[(Attrib.empty_binding, def1)] ||>
Sign.parent_path ||>>
@@ -224,7 +225,7 @@
Const (swap_name, swapT) $ x $ (prm $ xs $ a)));
in
thy' |>
- Primrec.add_primrec_global
+ Old_Primrec.add_primrec_global
[(Binding.name prm_name, SOME prmT, NoSyn)]
[(Attrib.empty_binding, def1), (Attrib.empty_binding, def2)] ||>
Sign.parent_path
--- a/src/HOL/Nominal/nominal_datatype.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML Mon Sep 01 16:17:46 2014 +0200
@@ -6,8 +6,9 @@
signature NOMINAL_DATATYPE =
sig
- val nominal_datatype : Datatype_Aux.config -> Datatype.spec list -> theory -> theory
- val nominal_datatype_cmd : Datatype_Aux.config -> Datatype.spec_cmd list -> theory -> theory
+ val nominal_datatype : Old_Datatype_Aux.config -> Old_Datatype.spec list -> theory -> theory
+ val nominal_datatype_cmd : Old_Datatype_Aux.config -> Old_Datatype.spec_cmd list -> theory ->
+ theory
type descr
type nominal_datatype_info
val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
@@ -42,8 +43,8 @@
(* theory data *)
type descr =
- (int * (string * Datatype_Aux.dtyp list *
- (string * (Datatype_Aux.dtyp list * Datatype_Aux.dtyp) list) list)) list;
+ (int * (string * Old_Datatype_Aux.dtyp list *
+ (string * (Old_Datatype_Aux.dtyp list * Old_Datatype_Aux.dtyp) list) list)) list;
type nominal_datatype_info =
{index : int,
@@ -83,7 +84,7 @@
(*******************************)
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of Datatype.distinct_lemma);
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of Old_Datatype.distinct_lemma);
(** simplification procedure for sorting permutations **)
@@ -199,10 +200,10 @@
val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
- val (full_new_type_names',thy1) = Datatype.add_datatype config dts'' thy;
+ val (full_new_type_names',thy1) = Old_Datatype.add_datatype config dts'' thy;
- val {descr, induct, ...} = Datatype_Data.the_info thy1 (hd full_new_type_names');
- fun nth_dtyp i = Datatype_Aux.typ_of_dtyp descr (Datatype_Aux.DtRec i);
+ val {descr, induct, ...} = Old_Datatype_Data.the_info thy1 (hd full_new_type_names');
+ fun nth_dtyp i = Old_Datatype_Aux.typ_of_dtyp descr (Old_Datatype_Aux.DtRec i);
val big_name = space_implode "_" new_type_names;
@@ -214,8 +215,8 @@
val perm_types = map (fn (i, _) =>
let val T = nth_dtyp i
in permT --> T --> T end) descr;
- val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) =>
- "perm_" ^ Datatype_Aux.name_of_typ (nth_dtyp i)) descr);
+ val perm_names' = Old_Datatype_Prop.indexify_names (map (fn (i, _) =>
+ "perm_" ^ Old_Datatype_Aux.name_of_typ (nth_dtyp i)) descr);
val perm_names = replicate (length new_type_names) @{const_name Nominal.perm} @
map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
val perm_names_types = perm_names ~~ perm_types;
@@ -225,17 +226,17 @@
let val T = nth_dtyp i
in map (fn (cname, dts) =>
let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr) dts;
- val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts;
+ val names = Name.variant_list ["pi"] (Old_Datatype_Prop.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 Datatype_Aux.is_rec_type dt then
+ in if Old_Datatype_Aux.is_rec_type dt then
let val Us = binder_types T
in
fold_rev (Term.abs o pair "x") Us
- (Free (nth perm_names_types' (Datatype_Aux.body_index dt)) $ pi $
+ (Free (nth perm_names_types' (Old_Datatype_Aux.body_index dt)) $ pi $
list_comb (x, map (fn (i, U) =>
Const (@{const_name Nominal.perm}, permT --> U --> U) $
(Const (@{const_name rev}, permT --> permT) $ pi) $
@@ -253,7 +254,7 @@
end) descr;
val (perm_simps, thy2) =
- Primrec.add_primrec_overloaded
+ Old_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;
@@ -264,12 +265,12 @@
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 = Datatype_Prop.make_tnames (map body_type perm_types);
+ val perm_indnames = Old_Datatype_Prop.make_tnames (map body_type perm_types);
val perm_fun_def = Simpdata.mk_eq @{thm perm_fun_def};
val unfolded_perm_eq_thms =
if length descr = length new_type_names then []
- else map Drule.export_without_context (List.drop (Datatype_Aux.split_conj_thm
+ else map Drule.export_without_context (List.drop (Old_Datatype_Aux.split_conj_thm
(Goal.prove_global_future thy2 [] []
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn (c as (s, T), x) =>
@@ -278,7 +279,7 @@
Const (@{const_name Nominal.perm}, T) $ pi $ Free (x, T2))
end)
(perm_names_types ~~ perm_indnames))))
- (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1,
+ (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac induct perm_indnames 1,
ALLGOALS (asm_full_simp_tac (ctxt addsimps [perm_fun_def]))])),
length new_type_names));
@@ -288,7 +289,7 @@
val perm_empty_thms = maps (fn a =>
let val permT = mk_permT (Type (a, []))
- in map Drule.export_without_context (List.take (Datatype_Aux.split_conj_thm
+ in map Drule.export_without_context (List.take (Old_Datatype_Aux.split_conj_thm
(Goal.prove_global_future thy2 [] []
(augment_sort thy2 [pt_class_of thy2 a]
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -298,7 +299,7 @@
Free (x, T)))
(perm_names ~~
map body_type perm_types ~~ perm_indnames)))))
- (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1,
+ (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac induct perm_indnames 1,
ALLGOALS (asm_full_simp_tac ctxt)])),
length new_type_names))
end)
@@ -320,7 +321,7 @@
val pt_inst = pt_inst_of thy2 a;
val pt2' = pt_inst RS pt2;
val pt2_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
- in List.take (map Drule.export_without_context (Datatype_Aux.split_conj_thm
+ in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm
(Goal.prove_global_future thy2 [] []
(augment_sort thy2 [pt_class_of thy2 a]
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -333,7 +334,7 @@
end)
(perm_names ~~
map body_type perm_types ~~ perm_indnames)))))
- (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1,
+ (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac induct perm_indnames 1,
ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt2', pt2_ax]))]))),
length new_type_names)
end) atoms;
@@ -355,7 +356,7 @@
val pt3' = pt_inst RS pt3;
val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
val pt3_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
- in List.take (map Drule.export_without_context (Datatype_Aux.split_conj_thm
+ in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm
(Goal.prove_global_future thy2 [] []
(augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
(HOLogic.mk_Trueprop (Const (@{const_name Nominal.prm_eq},
@@ -369,7 +370,7 @@
end)
(perm_names ~~
map body_type perm_types ~~ perm_indnames))))))
- (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1,
+ (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac induct perm_indnames 1,
ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt3', pt3_rev', pt3_ax]))]))),
length new_type_names)
end) atoms;
@@ -406,7 +407,7 @@
at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
end))
val sort = Sign.minimize_sort thy (Sign.certify_sort thy (cp_class :: pt_class));
- val thms = Datatype_Aux.split_conj_thm (Goal.prove_global_future thy [] []
+ val thms = Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy [] []
(augment_sort thy sort
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn ((s, T), x) =>
@@ -421,7 +422,7 @@
perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
end)
(perm_names ~~ Ts ~~ perm_indnames)))))
- (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1,
+ (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac induct perm_indnames 1,
ALLGOALS (asm_full_simp_tac (simps ctxt))]))
in
fold (fn (s, tvs) => fn thy => Axclass.prove_arity
@@ -459,25 +460,25 @@
val _ = warning "representing sets";
val rep_set_names =
- Datatype_Prop.indexify_names
- (map (fn (i, _) => Datatype_Aux.name_of_typ (nth_dtyp i) ^ "_set") descr);
+ Old_Datatype_Prop.indexify_names
+ (map (fn (i, _) => Old_Datatype_Aux.name_of_typ (nth_dtyp i) ^ "_set") descr);
val big_rep_name =
- space_implode "_" (Datatype_Prop.indexify_names (map_filter
+ space_implode "_" (Old_Datatype_Prop.indexify_names (map_filter
(fn (i, (@{type_name noption}, _, _)) => NONE
- | (i, _) => SOME (Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set";
+ | (i, _) => SOME (Old_Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set";
val _ = warning ("big_rep_name: " ^ big_rep_name);
- fun strip_option (dtf as Datatype_Aux.DtType ("fun", [dt, Datatype_Aux.DtRec i])) =
+ fun strip_option (dtf as Old_Datatype_Aux.DtType ("fun", [dt, Old_Datatype_Aux.DtRec i])) =
(case AList.lookup op = descr i of
SOME (@{type_name noption}, _, [(_, [dt']), _]) =>
apfst (cons dt) (strip_option dt')
| _ => ([], dtf))
- | strip_option (Datatype_Aux.DtType ("fun",
- [dt, Datatype_Aux.DtType (@{type_name noption}, [dt'])])) =
+ | strip_option (Old_Datatype_Aux.DtType ("fun",
+ [dt, Old_Datatype_Aux.DtType (@{type_name noption}, [dt'])])) =
apfst (cons dt) (strip_option dt')
| strip_option dt = ([], dt);
- val dt_atomTs = distinct op = (map (Datatype_Aux.typ_of_dtyp descr)
+ val dt_atomTs = distinct op = (map (Old_Datatype_Aux.typ_of_dtyp descr)
(maps (fn (_, (_, _, cs)) => maps (maps (fst o strip_option) o snd) cs) descr));
val dt_atoms = map (fst o dest_Type) dt_atomTs;
@@ -486,20 +487,20 @@
fun mk_prem dt (j, j', prems, ts) =
let
val (dts, dt') = strip_option dt;
- val (dts', dt'') = Datatype_Aux.strip_dtyp dt';
- val Ts = map (Datatype_Aux.typ_of_dtyp descr) dts;
- val Us = map (Datatype_Aux.typ_of_dtyp descr) dts';
- val T = Datatype_Aux.typ_of_dtyp descr dt'';
- val free = Datatype_Aux.mk_Free "x" (Us ---> T) j;
- val free' = Datatype_Aux.app_bnds free (length Us);
+ val (dts', dt'') = Old_Datatype_Aux.strip_dtyp dt';
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts;
+ val Us = map (Old_Datatype_Aux.typ_of_dtyp descr) dts';
+ val T = Old_Datatype_Aux.typ_of_dtyp descr dt'';
+ val free = Old_Datatype_Aux.mk_Free "x" (Us ---> T) j;
+ val free' = Old_Datatype_Aux.app_bnds free (length Us);
fun mk_abs_fun T (i, t) =
let val U = fastype_of t
in (i + 1, Const (@{const_name Nominal.abs_fun}, [T, U, T] --->
- Type (@{type_name noption}, [U])) $ Datatype_Aux.mk_Free "y" T i $ t)
+ Type (@{type_name noption}, [U])) $ Old_Datatype_Aux.mk_Free "y" T i $ t)
end
in (j + 1, j' + length Ts,
case dt'' of
- Datatype_Aux.DtRec k => Logic.list_all (map (pair "x") Us,
+ Old_Datatype_Aux.DtRec k => Logic.list_all (map (pair "x") Us,
HOLogic.mk_Trueprop (Free (nth rep_set_names k,
T --> HOLogic.boolT) $ free')) :: prems
| _ => prems,
@@ -545,7 +546,7 @@
(perm_indnames ~~ descr);
fun mk_perm_closed name = map (fn th => Drule.export_without_context (th RS mp))
- (List.take (Datatype_Aux.split_conj_thm (Goal.prove_global_future thy4 [] []
+ (List.take (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy4 [] []
(augment_sort thy4
(pt_class_of thy4 name :: map (cp_class_of thy4 name) (remove (op =) name dt_atoms))
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
@@ -558,7 +559,7 @@
Free ("pi", permT) $ Free (x, T)))
end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
(fn {context = ctxt, ...} => EVERY
- [Datatype_Aux.ind_tac rep_induct [] 1,
+ [Old_Datatype_Aux.ind_tac rep_induct [] 1,
ALLGOALS (simp_tac (ctxt addsimps
(Thm.symmetric perm_fun_def :: abs_perm))),
ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
@@ -680,8 +681,9 @@
(fn ((i, (@{type_name noption}, _, _)), p) => p
| ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
- fun reindex (Datatype_Aux.DtType (s, dts)) = Datatype_Aux.DtType (s, map reindex dts)
- | reindex (Datatype_Aux.DtRec i) = Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i))
+ fun reindex (Old_Datatype_Aux.DtType (s, dts)) = Old_Datatype_Aux.DtType (s, map reindex dts)
+ | reindex (Old_Datatype_Aux.DtRec i) =
+ Old_Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i))
| reindex dt = dt;
fun strip_suffix i s = implode (List.take (raw_explode s, size s - i)); (* FIXME Symbol.explode (?) *)
@@ -717,14 +719,14 @@
map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
(constrs ~~ idxss)))) (descr'' ~~ ndescr);
- fun nth_dtyp' i = Datatype_Aux.typ_of_dtyp descr'' (Datatype_Aux.DtRec i);
+ fun nth_dtyp' i = Old_Datatype_Aux.typ_of_dtyp descr'' (Old_Datatype_Aux.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 = Datatype_Aux.get_rec_types descr'';
+ val recTs = Old_Datatype_Aux.get_rec_types descr'';
val newTs' = take (length new_type_names) recTs';
val newTs = take (length new_type_names) recTs;
@@ -736,17 +738,20 @@
fun constr_arg (dts, dt) (j, l_args, r_args) =
let
val xs =
- map (fn (dt, i) => Datatype_Aux.mk_Free "x" (Datatype_Aux.typ_of_dtyp descr'' dt) i)
+ map (fn (dt, i) =>
+ Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) i)
(dts ~~ (j upto j + length dts - 1))
- val x = Datatype_Aux.mk_Free "x" (Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts)
+ val x =
+ Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
+ (j + length dts)
in
(j + length dts + 1,
xs @ x :: l_args,
fold_rev mk_abs_fun xs
(case dt of
- Datatype_Aux.DtRec k => if k < length new_type_names then
- Const (nth rep_names k, Datatype_Aux.typ_of_dtyp descr'' dt -->
- Datatype_Aux.typ_of_dtyp descr dt) $ x
+ Old_Datatype_Aux.DtRec k => if k < length new_type_names then
+ Const (nth rep_names k, Old_Datatype_Aux.typ_of_dtyp descr'' dt -->
+ Old_Datatype_Aux.typ_of_dtyp descr dt) $ x
else error "nested recursion not (yet) supported"
| _ => x) :: r_args)
end
@@ -773,7 +778,7 @@
(Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
val dist =
Drule.export_without_context
- (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] Datatype.distinct_lemma);
+ (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] Old_Datatype.distinct_lemma);
val (thy', defs', eqns') = fold (make_constr_def tname T T')
(constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, [])
in
@@ -829,7 +834,7 @@
(* prove distinctness theorems *)
- val distinct_props = Datatype_Prop.make_distincts descr';
+ val distinct_props = Old_Datatype_Prop.make_distincts descr';
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;
@@ -865,12 +870,13 @@
fun constr_arg (dts, dt) (j, l_args, r_args) =
let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr'') dts;
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') dts;
val xs =
- map (fn (T, i) => Datatype_Aux.mk_Free "x" T i)
+ map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i)
(Ts ~~ (j upto j + length dts - 1));
val x =
- Datatype_Aux.mk_Free "x" (Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts);
+ Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
+ (j + length dts);
in
(j + length dts + 1,
xs @ x :: l_args,
@@ -918,13 +924,15 @@
fun make_inj (dts, dt) (j, args1, args2, eqs) =
let
val Ts_idx =
- map (Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => Datatype_Aux.mk_Free "x" T i) Ts_idx;
- val ys = map (fn (T, i) => Datatype_Aux.mk_Free "y" T i) Ts_idx;
+ map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx;
+ val ys = map (fn (T, i) => Old_Datatype_Aux.mk_Free "y" T i) Ts_idx;
val x =
- Datatype_Aux.mk_Free "x" (Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts);
+ Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
+ (j + length dts);
val y =
- Datatype_Aux.mk_Free "y" (Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts);
+ Old_Datatype_Aux.mk_Free "y" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
+ (j + length dts);
in
(j + length dts + 1,
xs @ (x :: args1), ys @ (y :: args2),
@@ -965,10 +973,11 @@
fun process_constr (dts, dt) (j, args1, args2) =
let
val Ts_idx =
- map (Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => Datatype_Aux.mk_Free "x" T i) Ts_idx;
+ map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx;
val x =
- Datatype_Aux.mk_Free "x" (Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts);
+ Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
+ (j + length dts);
in
(j + length dts + 1,
xs @ (x :: args1), fold_rev mk_abs_fun xs x :: args2)
@@ -1007,16 +1016,17 @@
fun mk_indrule_lemma (((i, _), T), U) (prems, concls) =
let
- val Rep_t = Const (nth rep_names i, T --> U) $ Datatype_Aux.mk_Free "x" T i;
+ val Rep_t = Const (nth rep_names i, T --> U) $ Old_Datatype_Aux.mk_Free "x" T i;
val Abs_t = Const (nth abs_names i, U --> T);
in
(prems @ [HOLogic.imp $
(Const (nth rep_set_names'' i, U --> HOLogic.boolT) $ Rep_t) $
- (Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
+ (Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
concls @
- [Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ Datatype_Aux.mk_Free "x" T i])
+ [Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $
+ Old_Datatype_Aux.mk_Free "x" T i])
end;
val (indrule_lemma_prems, indrule_lemma_concls) =
@@ -1024,8 +1034,8 @@
val indrule_lemma = Goal.prove_global_future thy8 [] []
(Logic.mk_implies
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_concls)))
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
(fn {context = ctxt, ...} => EVERY
[REPEAT (etac conjE 1),
REPEAT (EVERY
@@ -1041,25 +1051,25 @@
val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
- val dt_induct_prop = Datatype_Prop.make_ind descr';
+ val dt_induct_prop = Old_Datatype_Prop.make_ind descr';
val dt_induct = Goal.prove_global_future thy8 []
(Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
(fn {prems, context = ctxt} => EVERY
[rtac indrule_lemma' 1,
- (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
+ (Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
EVERY (map (fn (prem, r) => (EVERY
[REPEAT (eresolve_tac Abs_inverse_thms' 1),
simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Thm.symmetric r]) 1,
DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
(prems ~~ constr_defs))]);
- val case_names_induct = Datatype_Data.mk_case_names_induct descr'';
+ val case_names_induct = Old_Datatype_Data.mk_case_names_induct descr'';
(**** prove that new datatypes have finite support ****)
val _ = warning "proving finite support for the new datatype";
- val indnames = Datatype_Prop.make_tnames recTs;
+ val indnames = Old_Datatype_Prop.make_tnames recTs;
val abs_supp = Global_Theory.get_thms thy8 "abs_supp";
val supp_atm = Global_Theory.get_thms thy8 "supp_atm";
@@ -1067,14 +1077,14 @@
val finite_supp_thms = map (fn atom =>
let val atomT = Type (atom, [])
in map Drule.export_without_context (List.take
- (Datatype_Aux.split_conj_thm (Goal.prove_global_future thy8 [] []
+ (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy8 [] []
(augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
(HOLogic.mk_Trueprop
(foldr1 HOLogic.mk_conj (map (fn (s, T) =>
Const (@{const_name finite}, HOLogic.mk_setT atomT --> HOLogic.boolT) $
(Const (@{const_name Nominal.supp}, T --> HOLogic.mk_setT atomT) $ Free (s, T)))
(indnames ~~ recTs)))))
- (fn {context = ctxt, ...} => Datatype_Aux.ind_tac dt_induct indnames 1 THEN
+ (fn {context = ctxt, ...} => Old_Datatype_Aux.ind_tac dt_induct indnames 1 THEN
ALLGOALS (asm_full_simp_tac (ctxt addsimps
(abs_supp @ supp_atm @
Global_Theory.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
@@ -1094,12 +1104,12 @@
Global_Theory.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
Global_Theory.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
Sign.parent_path ||>>
- Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
- Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
- Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
- Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>>
- Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>>
- Datatype_Aux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
+ Old_Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
+ Old_Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
+ Old_Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
+ Old_Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>>
+ Old_Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>>
+ Old_Datatype_Aux.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;
@@ -1119,7 +1129,7 @@
val fsT' = TFree ("'n", @{sort type});
val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
- (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
+ (Old_Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
fun make_pred fsT i T = Free (nth pnames i, fsT --> T --> HOLogic.boolT);
@@ -1137,11 +1147,11 @@
fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
let
- val recs = filter Datatype_Aux.is_rec_type cargs;
- val Ts = map (Datatype_Aux.typ_of_dtyp descr'') cargs;
- val recTs' = map (Datatype_Aux.typ_of_dtyp descr'') recs;
- val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts);
- val rec_tnames = map fst (filter (Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs));
+ val recs = filter Old_Datatype_Aux.is_rec_type cargs;
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs;
+ val recTs' = map (Old_Datatype_Aux.typ_of_dtyp descr'') recs;
+ val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts);
+ val rec_tnames = map fst (filter (Old_Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs));
val frees = tnames ~~ Ts;
val frees' = partition_cargs idxs frees;
val z = (singleton (Name.variant_list tnames) "z", fsT);
@@ -1153,8 +1163,8 @@
in
Logic.list_all (z :: map (pair "x") Us,
HOLogic.mk_Trueprop
- (make_pred fsT (Datatype_Aux.body_index dt) U $ Bound l $
- Datatype_Aux.app_bnds (Free (s, T)) l))
+ (make_pred fsT (Old_Datatype_Aux.body_index dt) U $ Bound l $
+ Old_Datatype_Aux.app_bnds (Free (s, T)) l))
end;
val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
@@ -1174,7 +1184,7 @@
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 = Datatype_Prop.make_tnames recTs;
+ val tnames = Old_Datatype_Prop.make_tnames recTs;
val zs = Name.variant_list tnames (replicate (length descr'') "z");
val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn ((((i, _), T), tname), z) =>
@@ -1198,7 +1208,7 @@
val induct' = Logic.list_implies (ind_prems', ind_concl');
val aux_ind_vars =
- (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~
+ (Old_Datatype_Prop.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 @{const_name HOL.conj})
@@ -1292,7 +1302,7 @@
val th = Goal.prove context [] []
(augment_sort thy9 fs_cp_sort aux_ind_concl)
(fn {context = context1, ...} =>
- EVERY (Datatype_Aux.ind_tac dt_induct tnames 1 ::
+ EVERY (Old_Datatype_Aux.ind_tac dt_induct tnames 1 ::
maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
map (fn ((cname, cargs), is) =>
REPEAT (rtac allI 1) THEN
@@ -1397,7 +1407,7 @@
val used = fold Term.add_tfree_namesT recTs [];
- val (rec_result_Ts', rec_fn_Ts') = Datatype_Prop.make_primrec_Ts descr' used;
+ val (rec_result_Ts', rec_fn_Ts') = Old_Datatype_Prop.make_primrec_Ts descr' used;
val rec_sort = if null dt_atomTs then @{sort type} else
Sign.minimize_sort thy10 (Sign.certify_sort thy10 pt_cp_sort);
@@ -1415,7 +1425,7 @@
(1 upto (length descr''));
val rec_set_names = map (Sign.full_bname thy10) rec_set_names';
- val rec_fns = map (uncurry (Datatype_Aux.mk_Free "f"))
+ val rec_fns = map (uncurry (Old_Datatype_Aux.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);
@@ -1440,13 +1450,13 @@
fun make_rec_intr T p rec_set ((cname, cargs), idxs)
(rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, l) =
let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr'') cargs;
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs;
val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
val frees' = partition_cargs idxs frees;
val binders = maps fst frees';
val atomTs = distinct op = (maps (map snd o fst) frees');
val recs = map_filter
- (fn ((_, Datatype_Aux.DtRec i), p) => SOME (i, p) | _ => NONE)
+ (fn ((_, Old_Datatype_Aux.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, _) => nth rec_result_Ts i) recs;
@@ -1508,7 +1518,7 @@
let
val permT = mk_permT aT;
val pi = Free ("pi", permT);
- val rec_fns_pi = map (mk_perm [] pi o uncurry (Datatype_Aux.mk_Free "f"))
+ val rec_fns_pi = map (mk_perm [] pi o uncurry (Old_Datatype_Aux.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);
@@ -1519,16 +1529,17 @@
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 => Drule.export_without_context (th RS mp)) (Datatype_Aux.split_conj_thm
- (Goal.prove_global_future thy11 [] []
- (augment_sort thy1 pt_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
- (fn {context = ctxt, ...} => rtac rec_induct 1 THEN REPEAT
- (simp_tac (put_simpset HOL_basic_ss ctxt
- addsimps flat perm_simps'
- addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- (resolve_tac rec_intrs THEN_ALL_NEW
- asm_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_bij @ perm_bij))) 1))))
+ val ths = map (fn th => Drule.export_without_context (th RS mp))
+ (Old_Datatype_Aux.split_conj_thm
+ (Goal.prove_global_future thy11 [] []
+ (augment_sort thy1 pt_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
+ (fn {context = ctxt, ...} => rtac rec_induct 1 THEN REPEAT
+ (simp_tac (put_simpset HOL_basic_ss ctxt
+ addsimps flat perm_simps'
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
+ (resolve_tac rec_intrs THEN_ALL_NEW
+ asm_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_bij @ perm_bij))) 1))))
val ths' = map (fn ((P, Q), th) =>
Goal.prove_global_future thy11 [] []
(augment_sort thy1 pt_cp_sort
@@ -1551,7 +1562,7 @@
(finite $ (Const (@{const_name Nominal.supp}, T --> aset) $ f)))
(rec_fns ~~ rec_fn_Ts)
in
- map (fn th => Drule.export_without_context (th RS mp)) (Datatype_Aux.split_conj_thm
+ map (fn th => Drule.export_without_context (th RS mp)) (Old_Datatype_Aux.split_conj_thm
(Goal.prove_global_future thy11 []
(map (augment_sort thy11 fs_cp_sort) fins)
(augment_sort thy11 fs_cp_sort
@@ -1643,10 +1654,10 @@
val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
val fun_tupleT = fastype_of fun_tuple;
val rec_unique_frees =
- Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs;
+ Old_Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs;
val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
val rec_unique_frees' =
- Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
+ Old_Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
val rec_unique_concls = map (fn ((x, U), R) =>
Const (@{const_name Ex1}, (U --> HOLogic.boolT) --> HOLogic.boolT) $
Abs ("y", U, R $ Free x $ Bound 0))
@@ -1688,7 +1699,7 @@
(Const (@{const_name finite}, HOLogic.mk_setT aT --> HOLogic.boolT) $
(Const (@{const_name Nominal.supp}, fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
- val rec_unique_thms = Datatype_Aux.split_conj_thm (Goal.prove
+ val rec_unique_thms = Old_Datatype_Aux.split_conj_thm (Goal.prove
(Proof_Context.init_global thy11) (map fst rec_unique_frees)
(map (augment_sort thy11 fs_cp_sort)
(flat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
@@ -1701,7 +1712,7 @@
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) (Datatype_Aux.split_conj_thm
+ val P_ths = map (fn th => th RS mp) (Old_Datatype_Aux.split_conj_thm
(Goal.prove context
(map fst (rec_unique_frees'' @ rec_unique_frees')) []
(augment_sort thy11 fs_cp_sort
@@ -2027,7 +2038,7 @@
resolve_tac rec_intrs 1,
REPEAT (solve (prems @ rec_total_thms) prems 1)])
end) (rec_eq_prems ~~
- Datatype_Prop.make_primrecs reccomb_names descr' thy12);
+ Old_Datatype_Prop.make_primrecs reccomb_names descr' thy12);
val dt_infos = map_index (make_dt_info pdescr induct reccomb_names rec_thms)
(descr1 ~~ distinct_thms ~~ inject_thms);
@@ -2047,12 +2058,12 @@
thy13
end;
-val nominal_datatype = gen_nominal_datatype Datatype.check_specs;
-val nominal_datatype_cmd = gen_nominal_datatype Datatype.read_specs;
+val nominal_datatype = gen_nominal_datatype Old_Datatype.check_specs;
+val nominal_datatype_cmd = gen_nominal_datatype Old_Datatype.read_specs;
val _ =
Outer_Syntax.command @{command_spec "nominal_datatype"} "define nominal datatypes"
- (Parse.and_list1 Datatype.spec_cmd >>
- (Toplevel.theory o nominal_datatype_cmd Datatype_Aux.default_config));
+ (Parse.and_list1 Old_Datatype.spec_cmd >>
+ (Toplevel.theory o nominal_datatype_cmd Old_Datatype_Aux.default_config));
end
--- a/src/HOL/Nominal/nominal_inductive.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML Mon Sep 01 16:17:46 2014 +0200
@@ -244,7 +244,7 @@
end) prems);
val ind_vars =
- (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~
+ (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~
map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
val ind_Ts = rev (map snd ind_vars);
@@ -645,7 +645,7 @@
val thss = map (fn atom =>
let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, [])))
in map (fn th => zero_var_indexes (th RS mp))
- (Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] []
+ (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] []
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p =>
let
val (h, ts) = strip_comb p;
--- a/src/HOL/Nominal/nominal_inductive2.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML Mon Sep 01 16:17:46 2014 +0200
@@ -261,7 +261,7 @@
in abs_params params' prem end) prems);
val ind_vars =
- (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~
+ (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~
map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
val ind_Ts = rev (map snd ind_vars);
--- a/src/HOL/Nominal/nominal_primrec.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Nominal/nominal_primrec.ML Mon Sep 01 16:17:46 2014 +0200
@@ -154,12 +154,12 @@
(fnames', fnss', (Const (@{const_name undefined}, dummyT))::fns))
| SOME (ls, cargs', rs, rhs, eq) =>
let
- val recs = filter (Datatype_Aux.is_rec_type o snd) (cargs' ~~ cargs);
+ val recs = filter (Old_Datatype_Aux.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, (Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
+ (Free x, (Old_Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
handle RecError s => primrec_eq_err lthy s eq
in (fnames'', fnss'', fold_rev absfree (cargs' @ subs) rhs' :: fns)
end)
@@ -190,7 +190,7 @@
NONE =>
let
val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
- replicate (length cargs + length (filter Datatype_Aux.is_rec_type cargs))
+ replicate (length cargs + length (filter Old_Datatype_Aux.is_rec_type cargs))
dummyT ---> HOLogic.unitT)) constrs;
val _ = warning ("No function definition for datatype " ^ quote tname)
in
--- a/src/HOL/Num.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Num.thy Mon Sep 01 16:17:46 2014 +0200
@@ -6,7 +6,7 @@
header {* Binary Numerals *}
theory Num
-imports Datatype BNF_LFP
+imports Old_Datatype BNF_LFP
begin
subsection {* The @{text num} type *}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Datatype.thy Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,523 @@
+(* Title: HOL/Old_Datatype.thy
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
+*)
+
+header {* Old Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums *}
+
+theory Old_Datatype
+imports Product_Type Sum_Type Nat
+keywords "datatype" :: thy_decl
+begin
+
+subsection {* The datatype universe *}
+
+definition "Node = {p. EX f x k. p = (f :: nat => 'b + nat, x ::'a + nat) & f k = Inr 0}"
+
+typedef ('a, 'b) node = "Node :: ((nat => 'b + nat) * ('a + nat)) set"
+ morphisms Rep_Node Abs_Node
+ unfolding Node_def by auto
+
+text{*Datatypes will be represented by sets of type @{text node}*}
+
+type_synonym 'a item = "('a, unit) node set"
+type_synonym ('a, 'b) dtree = "('a, 'b) node set"
+
+consts
+ Push :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))"
+
+ Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node"
+ ndepth :: "('a, 'b) node => nat"
+
+ Atom :: "('a + nat) => ('a, 'b) dtree"
+ Leaf :: "'a => ('a, 'b) dtree"
+ Numb :: "nat => ('a, 'b) dtree"
+ Scons :: "[('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree"
+ In0 :: "('a, 'b) dtree => ('a, 'b) dtree"
+ In1 :: "('a, 'b) dtree => ('a, 'b) dtree"
+ Lim :: "('b => ('a, 'b) dtree) => ('a, 'b) dtree"
+
+ ntrunc :: "[nat, ('a, 'b) dtree] => ('a, 'b) dtree"
+
+ uprod :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set"
+ usum :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set"
+
+ Split :: "[[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
+ Case :: "[[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
+
+ dprod :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set]
+ => (('a, 'b) dtree * ('a, 'b) dtree)set"
+ dsum :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set]
+ => (('a, 'b) dtree * ('a, 'b) dtree)set"
+
+
+defs
+
+ Push_Node_def: "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))"
+
+ (*crude "lists" of nats -- needed for the constructions*)
+ Push_def: "Push == (%b h. case_nat b h)"
+
+ (** operations on S-expressions -- sets of nodes **)
+
+ (*S-expression constructors*)
+ Atom_def: "Atom == (%x. {Abs_Node((%k. Inr 0, x))})"
+ Scons_def: "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)"
+
+ (*Leaf nodes, with arbitrary or nat labels*)
+ Leaf_def: "Leaf == Atom o Inl"
+ Numb_def: "Numb == Atom o Inr"
+
+ (*Injections of the "disjoint sum"*)
+ In0_def: "In0(M) == Scons (Numb 0) M"
+ In1_def: "In1(M) == Scons (Numb 1) M"
+
+ (*Function spaces*)
+ Lim_def: "Lim f == Union {z. ? x. z = Push_Node (Inl x) ` (f x)}"
+
+ (*the set of nodes with depth less than k*)
+ ndepth_def: "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)"
+ ntrunc_def: "ntrunc k N == {n. n:N & ndepth(n)<k}"
+
+ (*products and sums for the "universe"*)
+ uprod_def: "uprod A B == UN x:A. UN y:B. { Scons x y }"
+ usum_def: "usum A B == In0`A Un In1`B"
+
+ (*the corresponding eliminators*)
+ Split_def: "Split c M == THE u. EX x y. M = Scons x y & u = c x y"
+
+ Case_def: "Case c d M == THE u. (EX x . M = In0(x) & u = c(x))
+ | (EX y . M = In1(y) & u = d(y))"
+
+
+ (** equality for the "universe" **)
+
+ dprod_def: "dprod r s == UN (x,x'):r. UN (y,y'):s. {(Scons x y, Scons x' y')}"
+
+ dsum_def: "dsum r s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un
+ (UN (y,y'):s. {(In1(y),In1(y'))})"
+
+
+
+lemma apfst_convE:
+ "[| q = apfst f p; !!x y. [| p = (x,y); q = (f(x),y) |] ==> R
+ |] ==> R"
+by (force simp add: apfst_def)
+
+(** Push -- an injection, analogous to Cons on lists **)
+
+lemma Push_inject1: "Push i f = Push j g ==> i=j"
+apply (simp add: Push_def fun_eq_iff)
+apply (drule_tac x=0 in spec, simp)
+done
+
+lemma Push_inject2: "Push i f = Push j g ==> f=g"
+apply (auto simp add: Push_def fun_eq_iff)
+apply (drule_tac x="Suc x" in spec, simp)
+done
+
+lemma Push_inject:
+ "[| Push i f =Push j g; [| i=j; f=g |] ==> P |] ==> P"
+by (blast dest: Push_inject1 Push_inject2)
+
+lemma Push_neq_K0: "Push (Inr (Suc k)) f = (%z. Inr 0) ==> P"
+by (auto simp add: Push_def fun_eq_iff split: nat.split_asm)
+
+lemmas Abs_Node_inj = Abs_Node_inject [THEN [2] rev_iffD1]
+
+
+(*** Introduction rules for Node ***)
+
+lemma Node_K0_I: "(%k. Inr 0, a) : Node"
+by (simp add: Node_def)
+
+lemma Node_Push_I: "p: Node ==> apfst (Push i) p : Node"
+apply (simp add: Node_def Push_def)
+apply (fast intro!: apfst_conv nat.case(2)[THEN trans])
+done
+
+
+subsection{*Freeness: Distinctness of Constructors*}
+
+(** Scons vs Atom **)
+
+lemma Scons_not_Atom [iff]: "Scons M N \<noteq> Atom(a)"
+unfolding Atom_def Scons_def Push_Node_def One_nat_def
+by (blast intro: Node_K0_I Rep_Node [THEN Node_Push_I]
+ dest!: Abs_Node_inj
+ elim!: apfst_convE sym [THEN Push_neq_K0])
+
+lemmas Atom_not_Scons [iff] = Scons_not_Atom [THEN not_sym]
+
+
+(*** Injectiveness ***)
+
+(** Atomic nodes **)
+
+lemma inj_Atom: "inj(Atom)"
+apply (simp add: Atom_def)
+apply (blast intro!: inj_onI Node_K0_I dest!: Abs_Node_inj)
+done
+lemmas Atom_inject = inj_Atom [THEN injD]
+
+lemma Atom_Atom_eq [iff]: "(Atom(a)=Atom(b)) = (a=b)"
+by (blast dest!: Atom_inject)
+
+lemma inj_Leaf: "inj(Leaf)"
+apply (simp add: Leaf_def o_def)
+apply (rule inj_onI)
+apply (erule Atom_inject [THEN Inl_inject])
+done
+
+lemmas Leaf_inject [dest!] = inj_Leaf [THEN injD]
+
+lemma inj_Numb: "inj(Numb)"
+apply (simp add: Numb_def o_def)
+apply (rule inj_onI)
+apply (erule Atom_inject [THEN Inr_inject])
+done
+
+lemmas Numb_inject [dest!] = inj_Numb [THEN injD]
+
+
+(** Injectiveness of Push_Node **)
+
+lemma Push_Node_inject:
+ "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P
+ |] ==> P"
+apply (simp add: Push_Node_def)
+apply (erule Abs_Node_inj [THEN apfst_convE])
+apply (rule Rep_Node [THEN Node_Push_I])+
+apply (erule sym [THEN apfst_convE])
+apply (blast intro: Rep_Node_inject [THEN iffD1] trans sym elim!: Push_inject)
+done
+
+
+(** Injectiveness of Scons **)
+
+lemma Scons_inject_lemma1: "Scons M N <= Scons M' N' ==> M<=M'"
+unfolding Scons_def One_nat_def
+by (blast dest!: Push_Node_inject)
+
+lemma Scons_inject_lemma2: "Scons M N <= Scons M' N' ==> N<=N'"
+unfolding Scons_def One_nat_def
+by (blast dest!: Push_Node_inject)
+
+lemma Scons_inject1: "Scons M N = Scons M' N' ==> M=M'"
+apply (erule equalityE)
+apply (iprover intro: equalityI Scons_inject_lemma1)
+done
+
+lemma Scons_inject2: "Scons M N = Scons M' N' ==> N=N'"
+apply (erule equalityE)
+apply (iprover intro: equalityI Scons_inject_lemma2)
+done
+
+lemma Scons_inject:
+ "[| Scons M N = Scons M' N'; [| M=M'; N=N' |] ==> P |] ==> P"
+by (iprover dest: Scons_inject1 Scons_inject2)
+
+lemma Scons_Scons_eq [iff]: "(Scons M N = Scons M' N') = (M=M' & N=N')"
+by (blast elim!: Scons_inject)
+
+(*** Distinctness involving Leaf and Numb ***)
+
+(** Scons vs Leaf **)
+
+lemma Scons_not_Leaf [iff]: "Scons M N \<noteq> Leaf(a)"
+unfolding Leaf_def o_def by (rule Scons_not_Atom)
+
+lemmas Leaf_not_Scons [iff] = Scons_not_Leaf [THEN not_sym]
+
+(** Scons vs Numb **)
+
+lemma Scons_not_Numb [iff]: "Scons M N \<noteq> Numb(k)"
+unfolding Numb_def o_def by (rule Scons_not_Atom)
+
+lemmas Numb_not_Scons [iff] = Scons_not_Numb [THEN not_sym]
+
+
+(** Leaf vs Numb **)
+
+lemma Leaf_not_Numb [iff]: "Leaf(a) \<noteq> Numb(k)"
+by (simp add: Leaf_def Numb_def)
+
+lemmas Numb_not_Leaf [iff] = Leaf_not_Numb [THEN not_sym]
+
+
+(*** ndepth -- the depth of a node ***)
+
+lemma ndepth_K0: "ndepth (Abs_Node(%k. Inr 0, x)) = 0"
+by (simp add: ndepth_def Node_K0_I [THEN Abs_Node_inverse] Least_equality)
+
+lemma ndepth_Push_Node_aux:
+ "case_nat (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"
+apply (induct_tac "k", auto)
+apply (erule Least_le)
+done
+
+lemma ndepth_Push_Node:
+ "ndepth (Push_Node (Inr (Suc i)) n) = Suc(ndepth(n))"
+apply (insert Rep_Node [of n, unfolded Node_def])
+apply (auto simp add: ndepth_def Push_Node_def
+ Rep_Node [THEN Node_Push_I, THEN Abs_Node_inverse])
+apply (rule Least_equality)
+apply (auto simp add: Push_def ndepth_Push_Node_aux)
+apply (erule LeastI)
+done
+
+
+(*** ntrunc applied to the various node sets ***)
+
+lemma ntrunc_0 [simp]: "ntrunc 0 M = {}"
+by (simp add: ntrunc_def)
+
+lemma ntrunc_Atom [simp]: "ntrunc (Suc k) (Atom a) = Atom(a)"
+by (auto simp add: Atom_def ntrunc_def ndepth_K0)
+
+lemma ntrunc_Leaf [simp]: "ntrunc (Suc k) (Leaf a) = Leaf(a)"
+unfolding Leaf_def o_def by (rule ntrunc_Atom)
+
+lemma ntrunc_Numb [simp]: "ntrunc (Suc k) (Numb i) = Numb(i)"
+unfolding Numb_def o_def by (rule ntrunc_Atom)
+
+lemma ntrunc_Scons [simp]:
+ "ntrunc (Suc k) (Scons M N) = Scons (ntrunc k M) (ntrunc k N)"
+unfolding Scons_def ntrunc_def One_nat_def
+by (auto simp add: ndepth_Push_Node)
+
+
+
+(** Injection nodes **)
+
+lemma ntrunc_one_In0 [simp]: "ntrunc (Suc 0) (In0 M) = {}"
+apply (simp add: In0_def)
+apply (simp add: Scons_def)
+done
+
+lemma ntrunc_In0 [simp]: "ntrunc (Suc(Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"
+by (simp add: In0_def)
+
+lemma ntrunc_one_In1 [simp]: "ntrunc (Suc 0) (In1 M) = {}"
+apply (simp add: In1_def)
+apply (simp add: Scons_def)
+done
+
+lemma ntrunc_In1 [simp]: "ntrunc (Suc(Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"
+by (simp add: In1_def)
+
+
+subsection{*Set Constructions*}
+
+
+(*** Cartesian Product ***)
+
+lemma uprodI [intro!]: "[| M:A; N:B |] ==> Scons M N : uprod A B"
+by (simp add: uprod_def)
+
+(*The general elimination rule*)
+lemma uprodE [elim!]:
+ "[| c : uprod A B;
+ !!x y. [| x:A; y:B; c = Scons x y |] ==> P
+ |] ==> P"
+by (auto simp add: uprod_def)
+
+
+(*Elimination of a pair -- introduces no eigenvariables*)
+lemma uprodE2: "[| Scons M N : uprod A B; [| M:A; N:B |] ==> P |] ==> P"
+by (auto simp add: uprod_def)
+
+
+(*** Disjoint Sum ***)
+
+lemma usum_In0I [intro]: "M:A ==> In0(M) : usum A B"
+by (simp add: usum_def)
+
+lemma usum_In1I [intro]: "N:B ==> In1(N) : usum A B"
+by (simp add: usum_def)
+
+lemma usumE [elim!]:
+ "[| u : usum A B;
+ !!x. [| x:A; u=In0(x) |] ==> P;
+ !!y. [| y:B; u=In1(y) |] ==> P
+ |] ==> P"
+by (auto simp add: usum_def)
+
+
+(** Injection **)
+
+lemma In0_not_In1 [iff]: "In0(M) \<noteq> In1(N)"
+unfolding In0_def In1_def One_nat_def by auto
+
+lemmas In1_not_In0 [iff] = In0_not_In1 [THEN not_sym]
+
+lemma In0_inject: "In0(M) = In0(N) ==> M=N"
+by (simp add: In0_def)
+
+lemma In1_inject: "In1(M) = In1(N) ==> M=N"
+by (simp add: In1_def)
+
+lemma In0_eq [iff]: "(In0 M = In0 N) = (M=N)"
+by (blast dest!: In0_inject)
+
+lemma In1_eq [iff]: "(In1 M = In1 N) = (M=N)"
+by (blast dest!: In1_inject)
+
+lemma inj_In0: "inj In0"
+by (blast intro!: inj_onI)
+
+lemma inj_In1: "inj In1"
+by (blast intro!: inj_onI)
+
+
+(*** Function spaces ***)
+
+lemma Lim_inject: "Lim f = Lim g ==> f = g"
+apply (simp add: Lim_def)
+apply (rule ext)
+apply (blast elim!: Push_Node_inject)
+done
+
+
+(*** proving equality of sets and functions using ntrunc ***)
+
+lemma ntrunc_subsetI: "ntrunc k M <= M"
+by (auto simp add: ntrunc_def)
+
+lemma ntrunc_subsetD: "(!!k. ntrunc k M <= N) ==> M<=N"
+by (auto simp add: ntrunc_def)
+
+(*A generalized form of the take-lemma*)
+lemma ntrunc_equality: "(!!k. ntrunc k M = ntrunc k N) ==> M=N"
+apply (rule equalityI)
+apply (rule_tac [!] ntrunc_subsetD)
+apply (rule_tac [!] ntrunc_subsetI [THEN [2] subset_trans], auto)
+done
+
+lemma ntrunc_o_equality:
+ "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"
+apply (rule ntrunc_equality [THEN ext])
+apply (simp add: fun_eq_iff)
+done
+
+
+(*** Monotonicity ***)
+
+lemma uprod_mono: "[| A<=A'; B<=B' |] ==> uprod A B <= uprod A' B'"
+by (simp add: uprod_def, blast)
+
+lemma usum_mono: "[| A<=A'; B<=B' |] ==> usum A B <= usum A' B'"
+by (simp add: usum_def, blast)
+
+lemma Scons_mono: "[| M<=M'; N<=N' |] ==> Scons M N <= Scons M' N'"
+by (simp add: Scons_def, blast)
+
+lemma In0_mono: "M<=N ==> In0(M) <= In0(N)"
+by (simp add: In0_def Scons_mono)
+
+lemma In1_mono: "M<=N ==> In1(M) <= In1(N)"
+by (simp add: In1_def Scons_mono)
+
+
+(*** Split and Case ***)
+
+lemma Split [simp]: "Split c (Scons M N) = c M N"
+by (simp add: Split_def)
+
+lemma Case_In0 [simp]: "Case c d (In0 M) = c(M)"
+by (simp add: Case_def)
+
+lemma Case_In1 [simp]: "Case c d (In1 N) = d(N)"
+by (simp add: Case_def)
+
+
+
+(**** UN x. B(x) rules ****)
+
+lemma ntrunc_UN1: "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))"
+by (simp add: ntrunc_def, blast)
+
+lemma Scons_UN1_x: "Scons (UN x. f x) M = (UN x. Scons (f x) M)"
+by (simp add: Scons_def, blast)
+
+lemma Scons_UN1_y: "Scons M (UN x. f x) = (UN x. Scons M (f x))"
+by (simp add: Scons_def, blast)
+
+lemma In0_UN1: "In0(UN x. f(x)) = (UN x. In0(f(x)))"
+by (simp add: In0_def Scons_UN1_y)
+
+lemma In1_UN1: "In1(UN x. f(x)) = (UN x. In1(f(x)))"
+by (simp add: In1_def Scons_UN1_y)
+
+
+(*** Equality for Cartesian Product ***)
+
+lemma dprodI [intro!]:
+ "[| (M,M'):r; (N,N'):s |] ==> (Scons M N, Scons M' N') : dprod r s"
+by (auto simp add: dprod_def)
+
+(*The general elimination rule*)
+lemma dprodE [elim!]:
+ "[| c : dprod r s;
+ !!x y x' y'. [| (x,x') : r; (y,y') : s;
+ c = (Scons x y, Scons x' y') |] ==> P
+ |] ==> P"
+by (auto simp add: dprod_def)
+
+
+(*** Equality for Disjoint Sum ***)
+
+lemma dsum_In0I [intro]: "(M,M'):r ==> (In0(M), In0(M')) : dsum r s"
+by (auto simp add: dsum_def)
+
+lemma dsum_In1I [intro]: "(N,N'):s ==> (In1(N), In1(N')) : dsum r s"
+by (auto simp add: dsum_def)
+
+lemma dsumE [elim!]:
+ "[| w : dsum r s;
+ !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P;
+ !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P
+ |] ==> P"
+by (auto simp add: dsum_def)
+
+
+(*** Monotonicity ***)
+
+lemma dprod_mono: "[| r<=r'; s<=s' |] ==> dprod r s <= dprod r' s'"
+by blast
+
+lemma dsum_mono: "[| r<=r'; s<=s' |] ==> dsum r s <= dsum r' s'"
+by blast
+
+
+(*** Bounding theorems ***)
+
+lemma dprod_Sigma: "(dprod (A <*> B) (C <*> D)) <= (uprod A C) <*> (uprod B D)"
+by blast
+
+lemmas dprod_subset_Sigma = subset_trans [OF dprod_mono dprod_Sigma]
+
+(*Dependent version*)
+lemma dprod_subset_Sigma2:
+ "(dprod (Sigma A B) (Sigma C D)) <= Sigma (uprod A C) (Split (%x y. uprod (B x) (D y)))"
+by auto
+
+lemma dsum_Sigma: "(dsum (A <*> B) (C <*> D)) <= (usum A C) <*> (usum B D)"
+by blast
+
+lemmas dsum_subset_Sigma = subset_trans [OF dsum_mono dsum_Sigma]
+
+
+text {* hides popular names *}
+hide_type (open) node item
+hide_const (open) Push Node Atom Leaf Numb Lim Split Case
+
+ML_file "Tools/Old_Datatype/old_datatype.ML"
+
+ML_file "Tools/inductive_realizer.ML"
+setup InductiveRealizer.setup
+
+ML_file "Tools/Old_Datatype/old_datatype_realizer.ML"
+setup Old_Datatype_Realizer.setup
+
+end
--- a/src/HOL/Option.thy Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Option.thy Mon Sep 01 16:17:46 2014 +0200
@@ -5,7 +5,7 @@
header {* Datatype option *}
theory Option
-imports BNF_LFP Datatype Finite_Set
+imports BNF_LFP Old_Datatype Finite_Set
begin
datatype_new 'a option =
--- a/src/HOL/SPARK/Tools/spark_vcs.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Mon Sep 01 16:17:46 2014 +0200
@@ -173,8 +173,8 @@
fun add_enum_type tyname tyname' thy =
let
- val {case_name, ...} = the (Datatype_Data.get_info thy tyname');
- val cs = map Const (the (Datatype_Data.get_constrs thy tyname'));
+ val {case_name, ...} = the (Old_Datatype_Data.get_info thy tyname');
+ val cs = map Const (the (Old_Datatype_Data.get_constrs thy tyname'));
val k = length cs;
val T = Type (tyname', []);
val p = Const (@{const_name pos}, T --> HOLogic.intT);
@@ -209,7 +209,7 @@
(fn _ =>
rtac @{thm subset_antisym} 1 THEN
rtac @{thm subsetI} 1 THEN
- Datatype_Aux.exh_tac (K (#exhaust (Datatype_Data.the_info
+ Old_Datatype_Aux.exh_tac (K (#exhaust (Old_Datatype_Data.the_info
(Proof_Context.theory_of lthy) tyname'))) 1 THEN
ALLGOALS (asm_full_simp_tac lthy));
@@ -320,14 +320,14 @@
val tyname = Sign.full_name thy tyb
in
(thy |>
- Datatype.add_datatype {strict = true, quiet = true}
+ Old_Datatype.add_datatype {strict = true, quiet = true}
[((tyb, [], NoSyn),
map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
add_enum_type s tyname,
tyname)
end
| SOME (T as Type (tyname, []), cmap) =>
- (case Datatype_Data.get_constrs thy tyname of
+ (case Old_Datatype_Data.get_constrs thy tyname of
NONE => assoc_ty_err thy T s "is not a datatype"
| SOME cs =>
let val (prfx', _) = strip_prfx s
@@ -338,7 +338,7 @@
| SOME msg => assoc_ty_err thy T s msg
end)
| SOME (T, _) => assoc_ty_err thy T s "is not a datatype");
- val cs = map Const (the (Datatype_Data.get_constrs thy' tyname));
+ val cs = map Const (the (Old_Datatype_Data.get_constrs thy' tyname));
in
((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
fold Name.declare els ctxt),
@@ -888,7 +888,7 @@
handle Symtab.DUP _ => error ("SPARK type " ^ s ^
" already associated with type")) |>
(fn thy' =>
- case Datatype_Data.get_constrs thy' tyname of
+ case Old_Datatype_Data.get_constrs thy' tyname of
NONE => (case get_record_info thy' T of
NONE => thy'
| SOME {fields, ...} =>
--- a/src/HOL/Statespace/state_fun.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Statespace/state_fun.ML Mon Sep 01 16:17:46 2014 +0200
@@ -339,7 +339,7 @@
| mkName (TFree (x,_)) = mkUpper (Long_Name.base_name x)
| mkName (TVar ((x,_),_)) = mkUpper (Long_Name.base_name x);
-fun is_datatype thy = is_some o Datatype_Data.get_info thy;
+fun is_datatype thy = is_some o Old_Datatype_Data.get_info thy;
fun mk_map @{type_name List.list} = Syntax.const @{const_name List.map}
| mk_map n = Syntax.const ("StateFun.map_" ^ Long_Name.base_name n);
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Mon Sep 01 16:17:46 2014 +0200
@@ -1210,7 +1210,7 @@
is_some (fp_sugar_of no_defs_lthy bad_tc) then
error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
" in type expression " ^ fake_T_backdrop)
- else if is_some (Datatype_Data.get_info (Proof_Context.theory_of no_defs_lthy)
+ else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of no_defs_lthy)
bad_tc) then
error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
" via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Mon Sep 01 16:17:46 2014 +0200
@@ -329,7 +329,7 @@
fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
fun not_co_datatype (T as Type (s, _)) =
if fp = Least_FP andalso
- is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
+ is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
else
not_co_datatype0 T
--- a/src/HOL/Tools/BNF/bnf_lfp_compat.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML Mon Sep 01 16:17:46 2014 +0200
@@ -28,8 +28,9 @@
val kks = map fst desc;
val perm_kks = sort int_ord kks;
- fun perm_dtyp (Datatype_Aux.DtType (s, Ds)) = Datatype_Aux.DtType (s, map perm_dtyp Ds)
- | perm_dtyp (Datatype_Aux.DtRec kk) = Datatype_Aux.DtRec (find_index (curry (op =) kk) kks)
+ fun perm_dtyp (Old_Datatype_Aux.DtType (s, Ds)) = Old_Datatype_Aux.DtType (s, map perm_dtyp Ds)
+ | perm_dtyp (Old_Datatype_Aux.DtRec kk) =
+ Old_Datatype_Aux.DtRec (find_index (curry (op =) kk) kks)
| perm_dtyp D = D
in
if perm_kks = kks then
@@ -68,7 +69,7 @@
val nn_fp = length fpTs;
- val mk_dtyp = Datatype_Aux.dtyp_of_typ (map (apsnd (map Term.dest_TFree) o dest_Type) fpTs);
+ val mk_dtyp = Old_Datatype_Aux.dtyp_of_typ (map (apsnd (map Term.dest_TFree) o dest_Type) fpTs);
fun mk_ctr_descr Ts = mk_ctr Ts #> dest_Const ##> (binder_types #> map mk_dtyp);
fun mk_typ_descr index (Type (T_name, Ts)) ({ctrs, ...} : ctr_sugar) =
@@ -76,9 +77,9 @@
val fp_ctr_sugars = map (#ctr_sugar o lfp_sugar_of) fpT_names;
val orig_descr = map3 mk_typ_descr (0 upto nn_fp - 1) fpTs fp_ctr_sugars;
- val all_infos = Datatype_Data.get_all thy;
+ val all_infos = Old_Datatype_Data.get_all thy;
val (orig_descr' :: nested_descrs, _) =
- Datatype_Aux.unfold_datatypes lthy orig_descr all_infos orig_descr nn_fp;
+ Old_Datatype_Aux.unfold_datatypes lthy orig_descr all_infos orig_descr nn_fp;
fun cliquify_descr [] = []
| cliquify_descr [entry] = [[entry]]
@@ -90,7 +91,7 @@
else
(case Symtab.lookup all_infos T_name1 of
SOME {descr, ...} =>
- length (filter_out (exists Datatype_Aux.is_rec_type o #2 o snd) descr)
+ length (filter_out (exists Old_Datatype_Aux.is_rec_type o #2 o snd) descr)
| NONE => raise Fail "unknown old-style datatype");
in
chop nn full_descr ||> cliquify_descr |> op ::
@@ -102,15 +103,15 @@
split_list (flat (map_index (fn (i, descr) => map (pair i) descr)
(maps cliquify_descr descrs)));
- val dest_dtyp = Datatype_Aux.typ_of_dtyp descr;
+ val dest_dtyp = Old_Datatype_Aux.typ_of_dtyp descr;
- val Ts = Datatype_Aux.get_rec_types descr;
+ val Ts = Old_Datatype_Aux.get_rec_types descr;
val nn = length Ts;
val fp_sugars0 = map (lfp_sugar_of o fst o dest_Type) Ts;
val ctr_Tsss = map (map (map dest_dtyp o snd) o #3 o snd) descr;
val kkssss =
- map (map (map (fn Datatype_Aux.DtRec kk => [kk] | _ => []) o snd) o #3 o snd) descr;
+ map (map (map (fn Old_Datatype_Aux.DtRec kk => [kk] | _ => []) o snd) o #3 o snd) descr;
val callers = map (fn kk => Var ((Name.uu, kk), @{typ "unit => unit"})) (0 upto nn - 1);
@@ -175,8 +176,8 @@
end);
val register_interpret =
- Datatype_Data.register infos
- #> Datatype_Data.interpretation_data (Datatype_Aux.default_config, map fst infos)
+ Old_Datatype_Data.register infos
+ #> Old_Datatype_Data.interpretation_data (Old_Datatype_Aux.default_config, map fst infos)
in
lthy
|> Local_Theory.raw_theory register_interpret
--- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Mon Sep 01 16:17:46 2014 +0200
@@ -444,7 +444,7 @@
|> map (maps (map_filter (find_rec_calls has_call)));
fun is_only_old_datatype (Type (s, _)) =
- is_some (Datatype_Data.get_info thy s) andalso not (is_new_datatype lthy0 s)
+ is_some (Old_Datatype_Data.get_info thy s) andalso not (is_new_datatype lthy0 s)
| is_only_old_datatype _ = false;
val _ = if exists is_only_old_datatype arg_Ts then raise OLD_PRIMREC () else ();
@@ -561,8 +561,8 @@
end
handle OLD_PRIMREC () => old_primrec raw_fixes raw_specs lthy |>> apsnd single;
-val add_primrec = gen_primrec Primrec.add_primrec Specification.check_spec [];
-val add_primrec_cmd = gen_primrec Primrec.add_primrec_cmd Specification.read_spec;
+val add_primrec = gen_primrec Old_Primrec.add_primrec Specification.check_spec [];
+val add_primrec_cmd = gen_primrec Old_Primrec.add_primrec_cmd Specification.read_spec;
fun add_primrec_global fixes specs =
Named_Target.theory_init
--- a/src/HOL/Tools/Datatype/datatype.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,799 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package: definitional introduction of datatypes
-with proof of characteristic theorems: injectivity / distinctness
-of constructors and induction. Main interface to datatypes
-after full bootstrap of datatype package.
-*)
-
-signature DATATYPE =
-sig
- val distinct_lemma: thm
- type spec =
- (binding * (string * sort) list * mixfix) *
- (binding * typ list * mixfix) list
- type spec_cmd =
- (binding * (string * string option) list * mixfix) *
- (binding * string list * mixfix) list
- val read_specs: spec_cmd list -> theory -> spec list * Proof.context
- val check_specs: spec list -> theory -> spec list * Proof.context
- val add_datatype: Datatype_Aux.config -> spec list -> theory -> string list * theory
- val add_datatype_cmd: Datatype_Aux.config -> spec_cmd list -> theory -> string list * theory
- val spec_cmd: spec_cmd parser
-end;
-
-structure Datatype : DATATYPE =
-struct
-
-(** auxiliary **)
-
-val distinct_lemma = @{lemma "f x \<noteq> f y ==> x \<noteq> y" by iprover};
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
-
-fun exh_thm_of (dt_info : Datatype_Aux.info Symtab.table) tname =
- #exhaust (the (Symtab.lookup dt_info tname));
-
-val In0_inject = @{thm In0_inject};
-val In1_inject = @{thm In1_inject};
-val Scons_inject = @{thm Scons_inject};
-val Leaf_inject = @{thm Leaf_inject};
-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 Lim_inject = @{thm Lim_inject};
-val Inl_inject = @{thm Inl_inject};
-val Inr_inject = @{thm Inr_inject};
-val Suml_inject = @{thm Suml_inject};
-val Sumr_inject = @{thm Sumr_inject};
-
-val datatype_injI =
- @{lemma "(!!x. ALL y. f x = f y --> x = y) ==> inj f" by (simp add: inj_on_def)};
-
-
-(** proof of characteristic theorems **)
-
-fun representation_proofs (config : Datatype_Aux.config) (dt_info : Datatype_Aux.info Symtab.table)
- descr types_syntax constr_syntax case_names_induct thy =
- let
- val descr' = flat descr;
- val new_type_names = map (Binding.name_of o fst) types_syntax;
- val big_name = space_implode "_" new_type_names;
- val thy1 = Sign.add_path big_name thy;
- val big_rec_name = "rep_set_" ^ big_name;
- val rep_set_names' =
- if length descr' = 1 then [big_rec_name]
- else map (prefix (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 Datatype_Aux.dest_DtTFree Ts) (hd descr);
- val leafTs' = Datatype_Aux.get_nonrec_types descr';
- val branchTs = Datatype_Aux.get_branching_types descr';
- val branchT =
- if null branchTs then HOLogic.unitT
- else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) branchTs;
- val arities = remove (op =) 0 (Datatype_Aux.get_arities descr');
- val unneeded_vars =
- subtract (op =) (fold Term.add_tfreesT (leafTs' @ branchTs) []) (hd tyvars);
- val leafTs = leafTs' @ map TFree unneeded_vars;
- val recTs = Datatype_Aux.get_rec_types descr';
- val (newTs, oldTs) = chop (length (hd descr)) recTs;
- val sumT =
- if null leafTs then HOLogic.unitT
- else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) leafTs;
- val Univ_elT = HOLogic.mk_setT (Type (@{type_name Datatype.node}, [sumT, branchT]));
- val UnivT = HOLogic.mk_setT Univ_elT;
- val UnivT' = Univ_elT --> HOLogic.boolT;
- val Collect = Const (@{const_name Collect}, UnivT' --> UnivT);
-
- val In0 = Const (@{const_name Datatype.In0}, Univ_elT --> Univ_elT);
- val In1 = Const (@{const_name Datatype.In1}, Univ_elT --> Univ_elT);
- val Leaf = Const (@{const_name Datatype.Leaf}, sumT --> Univ_elT);
- val Lim = Const (@{const_name Datatype.Lim}, (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 (@{const_name Inl}, T1 --> T) $ mk_inj' T1 n2 i
- else Const (@{const_name Inr}, T2 --> T) $ mk_inj' T2 (n - n2) (i - n2)
- end;
- in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) end;
-
- (* make injections for constructors *)
-
- fun mk_univ_inj ts = Balanced_Tree.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 @{const_name Datatype.Scons}) 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 (@{const_name Sum_Type.Suml}, mkT T1) $ mk_inj T1 n2 i
- else Const (@{const_name Sum_Type.Sumr}, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
- end;
- in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) end;
-
- fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t;
-
- (************** generate introduction rules for representing set **********)
-
- val _ = Datatype_Aux.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 Datatype_Aux.strip_dtyp dt of
- (dts, Datatype_Aux.DtRec k) =>
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') dts;
- val free_t =
- Datatype_Aux.app_bnds (Datatype_Aux.mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
- in
- (j + 1, Logic.list_all (map (pair "x") Ts,
- HOLogic.mk_Trueprop
- (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
- mk_lim free_t Ts :: ts)
- end
- | _ =>
- let val T = Datatype_Aux.typ_of_dtyp descr' dt
- in (j + 1, prems, (Leaf $ mk_inj T (Datatype_Aux.mk_Free "x" T j)) :: ts) end);
-
- val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
- 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) =
- thy1
- |> Sign.map_naming Name_Space.conceal
- |> Inductive.add_inductive_global
- {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
- coind = false, no_elim = true, no_ind = false, skip_mono = true}
- (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
- (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
- ||> Sign.restore_naming thy1;
-
- (********************************* typedef ********************************)
-
- val (typedefs, thy3) = thy2
- |> Sign.parent_path
- |> fold_map
- (fn (((name, mx), tvs), c) =>
- Typedef.add_typedef_global (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 ~~ take (length newTs) rep_set_names)
- ||> Sign.add_path big_name;
-
- (*********************** definition of constructors ***********************)
-
- val big_rep_name = big_name ^ "_Rep_";
- val rep_names' = map (fn i => big_rep_name ^ string_of_int i) (1 upto length (flat (tl descr)));
- val all_rep_names =
- map (#Rep_name o #1 o #2) typedefs @
- 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 (typedef: Typedef.info) T n
- ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
- let
- fun constr_arg dt (j, l_args, r_args) =
- let
- val T = Datatype_Aux.typ_of_dtyp descr' dt;
- val free_t = Datatype_Aux.mk_Free "x" T j;
- in
- (case (Datatype_Aux.strip_dtyp dt, strip_type T) of
- ((_, Datatype_Aux.DtRec m), (Us, U)) =>
- (j + 1, free_t :: l_args, mk_lim
- (Const (nth all_rep_names m, U --> Univ_elT) $
- Datatype_Aux.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) = fold_rev constr_arg cargs (1, [], []);
- val constrT = map (Datatype_Aux.typ_of_dtyp descr') cargs ---> T;
- val ({Abs_name, Rep_name, ...}, _) = typedef;
- 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 = Thm.def_name (Long_Name.base_name cname);
- val eqn =
- HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (Rep_name, T --> Univ_elT) $ lhs, rhs));
- val ([def_thm], thy') =
- thy
- |> Sign.add_consts [(cname', constrT, mx)]
- |> (Global_Theory.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 (((((_, (_, _, constrs)), tname), typedef: Typedef.info), T), constr_syntax)
- (thy, defs, eqns, rep_congs, dist_lemmas) =
- let
- val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
- val rep_const = cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
- val cong' = cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong;
- val dist = cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma;
- val (thy', defs', eqns', _) =
- fold (make_constr_def typedef T (length constrs))
- (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
- in
- (Sign.parent_path thy', defs', eqns @ [eqns'],
- rep_congs @ [cong'], dist_lemmas @ [dist])
- end;
-
- val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
- fold dt_constr_defs
- (hd descr ~~ new_type_names ~~ map #2 typedefs ~~ newTs ~~ constr_syntax)
- (thy3 |> Sign.add_consts iso_decls |> Sign.parent_path, [], [], [], []);
-
-
- (*********** isomorphisms for new types (introduced by typedef) ***********)
-
- val _ = Datatype_Aux.message config "Proving isomorphism properties ...";
-
- val collect_simp = rewrite_rule (Proof_Context.init_global thy4) [mk_meta_eq mem_Collect_eq];
-
- val newT_iso_axms = typedefs |> map (fn (_, (_, {Abs_inverse, Rep_inverse, Rep, ...})) =>
- (collect_simp Abs_inverse, Rep_inverse, collect_simp Rep));
-
- val newT_iso_inj_thms = typedefs |> map (fn (_, (_, {Abs_inject, Rep_inject, ...})) =>
- (collect_simp Abs_inject RS iffD1, Rep_inject RS iffD1));
-
- (********* isomorphisms between existing types and "unfolded" types *******)
-
- (*---------------------------------------------------------------------*)
- (* isomorphisms are defined using primrec-combinators: *)
- (* generate appropriate functions for instantiating primrec-combinator *)
- (* *)
- (* e.g. Rep_dt_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
- (* *)
- (* also generate characteristic equations for isomorphisms *)
- (* *)
- (* e.g. Rep_dt_i (cons h t) = In1 (Scons (Rep_dt_j h) (Rep_dt_i t)) *)
- (*---------------------------------------------------------------------*)
-
- fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
- let
- val argTs = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val T = nth recTs k;
- val rep_const = Const (nth all_rep_names k, T --> Univ_elT);
- val constr = Const (cname, argTs ---> T);
-
- fun process_arg ks' dt (i2, i2', ts, Ts) =
- let
- val T' = Datatype_Aux.typ_of_dtyp descr' dt;
- val (Us, U) = strip_type T'
- in
- (case Datatype_Aux.strip_dtyp dt of
- (_, Datatype_Aux.DtRec j) =>
- if member (op =) ks' j then
- (i2 + 1, i2' + 1, ts @ [mk_lim (Datatype_Aux.app_bnds
- (Datatype_Aux.mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
- Ts @ [Us ---> Univ_elT])
- else
- (i2 + 1, i2', ts @ [mk_lim
- (Const (nth all_rep_names j, U --> Univ_elT) $
- Datatype_Aux.app_bnds (Datatype_Aux.mk_Free "x" T' i2) (length Us)) Us], Ts)
- | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (Datatype_Aux.mk_Free "x" T' i2)], Ts))
- end;
-
- val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
- val xs = map (uncurry (Datatype_Aux.mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
- val ys = map (uncurry (Datatype_Aux.mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
- val f = fold_rev lambda (xs @ ys) (mk_univ_inj ts n i);
-
- val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
- 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 (k, (_, _, constrs)) (fs, eqns, isos) =
- let
- val (fs', eqns', _) = fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
- val iso = (nth recTs k, nth all_rep_names k);
- in (fs', eqns', isos @ [iso]) end;
-
- val (fs, eqns, isos) = fold process_dt ds ([], [], []);
- val fTs = map fastype_of fs;
- val defs =
- map (fn (rec_name, (T, iso_name)) =>
- (Binding.name (Thm.def_name (Long_Name.base_name iso_name)),
- 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') =
- (Global_Theory.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 => Goal.prove_sorry_global thy' [] [] eqn
- (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt rewrites, rtac refl 1])) eqns;
-
- in (thy', char_thms' @ char_thms) end;
-
- val (thy5, iso_char_thms) =
- fold_rev make_iso_defs (tl descr) (Sign.add_path big_name thy4, []);
-
- (* 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.legacy_freeze prop;
- val used = Term.add_tfree_names a [];
-
- fun mk_thm i =
- let
- val Ts = map (TFree o rpair @{sort type}) (Name.variant_list used (replicate i "'t"));
- val f = Free ("f", Ts ---> U);
- in
- Goal.prove_sorry_global thy [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.list_all
- (map (pair "x") Ts, S $ Datatype_Aux.app_bnds f i)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (Term.abs o pair "x") Ts
- (r $ (a $ Datatype_Aux.app_bnds f i)), f))))
- (fn _ => EVERY [REPEAT_DETERM_N i (rtac @{thm 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 Rep_dt_i and Rep_dt_i x : rep_set_dt_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 induct = #induct (the (Symtab.lookup dt_info tname));
-
- fun mk_ind_concl (i, _) =
- let
- val T = nth recTs i;
- val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
- val rep_set_name = nth rep_set_names i;
- val concl1 =
- HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
- HOLogic.mk_eq (Rep_t $ Datatype_Aux.mk_Free "x" T i, Rep_t $ Bound 0) $
- HOLogic.mk_eq (Datatype_Aux.mk_Free "x" T i, Bound 0));
- val concl2 = Const (rep_set_name, UnivT') $ (Rep_t $ Datatype_Aux.mk_Free "x" T i);
- in (concl1, concl2) end;
-
- val (ind_concl1, ind_concl2) = split_list (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 =
- Goal.prove_sorry_global thy5 [] []
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl1))
- (fn {context = ctxt, ...} => EVERY
- [(Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
- REPEAT (EVERY
- [rtac allI 1, rtac impI 1,
- Datatype_Aux.exh_tac (exh_thm_of dt_info) 1,
- REPEAT (EVERY
- [hyp_subst_tac ctxt 1,
- rewrite_goals_tac ctxt 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 @{thm 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 datatype_injI) (Datatype_Aux.split_conj_thm inj_thm);
-
- val elem_thm =
- Goal.prove_sorry_global thy5 [] []
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl2))
- (fn {context = ctxt, ...} =>
- EVERY [
- (Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
- rewrite_goals_tac ctxt 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 @ Datatype_Aux.split_conj_thm elem_thm) end;
-
- val (iso_inj_thms_unfolded, iso_elem_thms) =
- fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
- val iso_inj_thms =
- map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
-
- (* prove rep_set_dt_i x --> x : range Rep_dt_i *)
-
- fun mk_iso_t (((set_name, iso_name), i), T) =
- let val isoT = T --> Univ_elT in
- HOLogic.imp $
- (Const (set_name, UnivT') $ Datatype_Aux.mk_Free "x" Univ_elT i) $
- (if i < length newTs then @{term True}
- else HOLogic.mk_mem (Datatype_Aux.mk_Free "x" Univ_elT i,
- Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
- Const (iso_name, isoT) $ Const (@{const_abbrev UNIV}, HOLogic.mk_setT T)))
- end;
-
- val iso_t = HOLogic.mk_Trueprop (Datatype_Aux.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
- drop (length newTs) (Datatype_Aux.split_conj_thm
- (Goal.prove_sorry_global thy5 [] [] iso_t (fn {context = ctxt, ...} => EVERY
- [(Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
- REPEAT (rtac TrueI 1),
- rewrite_goals_tac ctxt (mk_meta_eq @{thm choice_eq} ::
- Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
- rewrite_goals_tac ctxt (map Thm.symmetric range_eqs),
- REPEAT (EVERY
- [REPEAT (eresolve_tac ([rangeE, @{thm ex1_implies_ex} RS exE] @
- maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
- TRY (hyp_subst_tac ctxt 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 => @{thm f_the_inv_into_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 _ = Datatype_Aux.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
- Goal.prove_sorry_global thy5 [] [] eqn
- (fn {context = ctxt, ...} => EVERY
- [resolve_tac inj_thms 1,
- rewrite_goals_tac ctxt 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 dist_rewrites' =
- let
- fun prove [] = []
- | prove (t :: ts) =
- let
- val dist_thm = Goal.prove_sorry_global thy5 [] [] t (fn {context = ctxt, ...} =>
- EVERY [simp_tac (put_simpset HOL_ss ctxt addsimps dist_rewrites') 1])
- in dist_thm :: Drule.zero_var_indexes (dist_thm RS not_sym) :: prove ts end;
- in prove end;
-
- val distinct_thms =
- map2 (prove_distinct_thms) dist_rewrites (Datatype_Prop.make_distincts descr);
-
- (* 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
- Goal.prove_sorry_global thy5 [] [] t
- (fn {context = ctxt, ...} => EVERY
- [rtac iffI 1,
- REPEAT (etac conjE 2), hyp_subst_tac ctxt 2, rtac refl 2,
- dresolve_tac rep_congs 1, dtac @{thm box_equals} 1,
- REPEAT (resolve_tac rep_thms 1),
- REPEAT (eresolve_tac inj_thms 1),
- REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac @{thm 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)
- (Datatype_Prop.make_injs descr ~~ constr_rep_thms);
-
- val ((constr_inject', distinct_thms'), thy6) =
- thy5
- |> Sign.parent_path
- |> Datatype_Aux.store_thmss "inject" new_type_names constr_inject
- ||>> Datatype_Aux.store_thmss "distinct" new_type_names distinct_thms;
-
- (*************************** induction theorem ****************************)
-
- val _ = Datatype_Aux.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 @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded;
- val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
-
- fun mk_indrule_lemma (i, _) T =
- let
- val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ Datatype_Aux.mk_Free "x" T i;
- val Abs_t =
- if i < length newTs then
- Const (#Abs_name (#1 (#2 (nth typedefs i))), Univ_elT --> T)
- else
- Const (@{const_name the_inv_into},
- [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
- HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT);
- val prem =
- HOLogic.imp $
- (Const (nth rep_set_names i, UnivT') $ Rep_t) $
- (Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t));
- val concl =
- Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ Datatype_Aux.mk_Free "x" T i;
- in (prem, concl) end;
-
- val (indrule_lemma_prems, indrule_lemma_concls) =
- split_list (map2 mk_indrule_lemma descr' recTs);
-
- val cert = cterm_of thy6;
-
- val indrule_lemma =
- Goal.prove_sorry_global thy6 [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (Datatype_Aux.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 = Datatype_Prop.make_ind descr;
- val dt_induct =
- Goal.prove_sorry_global thy6 []
- (Logic.strip_imp_prems dt_induct_prop)
- (Logic.strip_imp_concl dt_induct_prop)
- (fn {context = ctxt, prems, ...} =>
- EVERY
- [rtac indrule_lemma' 1,
- (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
- EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms 1),
- simp_tac (put_simpset HOL_basic_ss ctxt
- addsimps (Thm.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
- |> Global_Theory.note_thmss ""
- [((Binding.qualify true big_name (Binding.name "induct"), [case_names_induct]),
- [([dt_induct], [])])];
- in
- ((constr_inject', distinct_thms', dt_induct'), thy7)
- end;
-
-
-
-(** datatype definition **)
-
-(* specifications *)
-
-type spec = (binding * (string * sort) list * mixfix) * (binding * typ list * mixfix) list;
-
-type spec_cmd =
- (binding * (string * string option) list * mixfix) * (binding * string list * mixfix) list;
-
-local
-
-fun parse_spec ctxt ((b, args, mx), constrs) =
- ((b, map (apsnd (Typedecl.read_constraint ctxt)) args, mx),
- constrs |> map (fn (c, Ts, mx') => (c, map (Syntax.parse_typ ctxt) Ts, mx')));
-
-fun check_specs ctxt (specs: spec list) =
- let
- fun prep_spec ((tname, args, mx), constrs) tys =
- let
- val (args', tys1) = chop (length args) tys;
- val (constrs', tys3) = (constrs, tys1) |-> fold_map (fn (cname, cargs, mx') => fn tys2 =>
- let val (cargs', tys3) = chop (length cargs) tys2;
- in ((cname, cargs', mx'), tys3) end);
- in (((tname, map dest_TFree args', mx), constrs'), tys3) end;
-
- val all_tys =
- specs |> maps (fn ((_, args, _), cs) => map TFree args @ maps #2 cs)
- |> Syntax.check_typs ctxt;
-
- in #1 (fold_map prep_spec specs all_tys) end;
-
-fun prep_specs parse raw_specs thy =
- let
- val ctxt = thy
- |> Sign.add_types_global (map (fn ((b, args, mx), _) => (b, length args, mx)) raw_specs)
- |> Proof_Context.init_global
- |> fold (fn ((_, args, _), _) => fold (fn (a, _) =>
- Variable.declare_typ (TFree (a, dummyS))) args) raw_specs;
- val specs = check_specs ctxt (map (parse ctxt) raw_specs);
- in (specs, ctxt) end;
-
-in
-
-val read_specs = prep_specs parse_spec;
-val check_specs = prep_specs (K I);
-
-end;
-
-
-(* main commands *)
-
-fun gen_add_datatype prep_specs config raw_specs thy =
- let
- val _ = Theory.requires thy (Context.theory_name @{theory}) "datatype definitions";
-
- val (dts, spec_ctxt) = prep_specs raw_specs thy;
- val ((_, tyvars, _), _) :: _ = dts;
- val string_of_tyvar = Syntax.string_of_typ spec_ctxt o TFree;
-
- val (new_dts, types_syntax) = dts |> map (fn ((tname, tvs, mx), _) =>
- let val full_tname = Sign.full_name thy tname in
- (case duplicates (op =) tvs of
- [] =>
- if eq_set (op =) (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 " ^ Binding.print tname ^
- " : " ^ commas (map string_of_tyvar dups)))
- end) |> split_list;
- val dt_names = map fst new_dts;
-
- val _ =
- (case duplicates (op =) (map fst new_dts) of
- [] => ()
- | dups => error ("Duplicate datatypes: " ^ commas_quote dups));
-
- fun prep_dt_spec ((tname, tvs, _), constrs) (dts', constr_syntax, i) =
- let
- fun prep_constr (cname, cargs, mx) (constrs, constr_syntax') =
- let
- val _ =
- (case subtract (op =) tvs (fold Term.add_tfreesT cargs []) of
- [] => ()
- | vs => error ("Extra type variables on rhs: " ^ commas (map string_of_tyvar vs)));
- val c = Sign.full_name_path thy (Binding.name_of tname) cname;
- in
- (constrs @ [(c, map (Datatype_Aux.dtyp_of_typ new_dts) cargs)],
- constr_syntax' @ [(cname, mx)])
- end handle ERROR msg =>
- cat_error msg ("The error above occurred in constructor " ^ Binding.print cname ^
- " of datatype " ^ Binding.print tname);
-
- val (constrs', constr_syntax') = fold prep_constr constrs ([], []);
- in
- (case duplicates (op =) (map fst constrs') of
- [] =>
- (dts' @ [(i, (Sign.full_name thy tname, map Datatype_Aux.DtTFree tvs, constrs'))],
- constr_syntax @ [constr_syntax'], i + 1)
- | dups =>
- error ("Duplicate constructors " ^ commas_quote dups ^
- " in datatype " ^ Binding.print tname))
- end;
-
- val (dts', constr_syntax, i) = fold prep_dt_spec dts ([], [], 0);
-
- val dt_info = Datatype_Data.get_all thy;
- val (descr, _) = Datatype_Aux.unfold_datatypes spec_ctxt dts' dt_info dts' i;
- val _ =
- Datatype_Aux.check_nonempty descr
- handle (exn as Datatype_Aux.Datatype_Empty s) =>
- if #strict config then error ("Nonemptiness check failed for datatype " ^ quote s)
- else reraise exn;
-
- val _ =
- Datatype_Aux.message config
- ("Constructing datatype(s) " ^ commas_quote (map (Binding.name_of o #1 o #1) dts));
- in
- thy
- |> representation_proofs config dt_info descr types_syntax constr_syntax
- (Datatype_Data.mk_case_names_induct (flat descr))
- |-> (fn (inject, distinct, induct) =>
- Rep_Datatype.derive_datatype_props config dt_names descr induct inject distinct)
- end;
-
-val add_datatype = gen_add_datatype check_specs;
-val add_datatype_cmd = gen_add_datatype read_specs;
-
-
-(* outer syntax *)
-
-val spec_cmd =
- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
- (@{keyword "="} |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix))
- >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Parse.triple1 cons));
-
-val _ =
- Outer_Syntax.command @{command_spec "datatype"} "define inductive datatypes"
- (Parse.and_list1 spec_cmd
- >> (Toplevel.theory o (snd oo add_datatype_cmd Datatype_Aux.default_config)));
-
-
-open Datatype_Data;
-
-end;
--- a/src/HOL/Tools/Datatype/datatype_aux.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,402 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype_aux.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package: auxiliary data structures and functions.
-*)
-
-signature DATATYPE_COMMON =
-sig
- type config = {strict : bool, quiet : bool}
- val default_config : config
- datatype dtyp =
- DtTFree of string * sort
- | DtType of string * dtyp list
- | DtRec of int
- type descr = (int * (string * dtyp list * (string * dtyp list) list)) list
- type info =
- {index : int,
- descr : descr,
- inject : thm list,
- distinct : thm list,
- induct : thm,
- inducts : thm list,
- exhaust : thm,
- nchotomy : thm,
- rec_names : string list,
- rec_rewrites : thm list,
- case_name : string,
- case_rewrites : thm list,
- case_cong : thm,
- case_cong_weak : thm,
- split : thm,
- split_asm: thm}
-end
-
-signature DATATYPE_AUX =
-sig
- include DATATYPE_COMMON
-
- val message : config -> string -> unit
-
- 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 ind_tac : thm -> string list -> int -> tactic
- val exh_tac : (string -> thm) -> int -> tactic
-
- exception Datatype
- exception Datatype_Empty of string
- val name_of_typ : typ -> string
- val dtyp_of_typ : (string * (string * sort) list) list -> typ -> dtyp
- val mk_Free : string -> typ -> int -> term
- val is_rec_type : dtyp -> bool
- val typ_of_dtyp : descr -> dtyp -> typ
- val dest_DtTFree : dtyp -> string * sort
- 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 -> typ list
- val get_branching_types : descr -> typ list
- val get_arities : descr -> int list
- val get_rec_types : descr -> typ list
- val interpret_construction : descr -> (string * sort) list ->
- {atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a} ->
- ((string * typ list) * (string * 'a list) list) list
- val check_nonempty : descr list -> unit
- val unfold_datatypes : Proof.context -> descr -> info Symtab.table ->
- descr -> int -> descr list * int
- val find_shortest_path : descr -> int -> (string * int) option
-end;
-
-structure Datatype_Aux : DATATYPE_AUX =
-struct
-
-(* datatype option flags *)
-
-type config = {strict : bool, quiet : bool};
-val default_config : config = {strict = true, quiet = false};
-
-fun message ({quiet = true, ...} : config) s = writeln s
- | message _ _ = ();
-
-
-(* store theorems in theory *)
-
-fun store_thmss_atts name tnames attss thmss =
- fold_map (fn ((tname, atts), thms) =>
- Global_Theory.note_thmss ""
- [((Binding.qualify true tname (Binding.name name), atts), [(thms, [])])]
- #-> (fn [(_, res)] => pair res)) (tnames ~~ attss ~~ thmss);
-
-fun store_thmss name tnames = store_thmss_atts name tnames (replicate (length tnames) []);
-
-fun store_thms_atts name tnames attss thms =
- fold_map (fn ((tname, atts), thm) =>
- Global_Theory.note_thmss ""
- [((Binding.qualify true tname (Binding.name name), atts), [([thm], [])])]
- #-> (fn [(_, [res])] => pair res)) (tnames ~~ attss ~~ thms);
-
-fun store_thms name tnames = store_thms_atts name 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 @{const_name HOL.conj});
-val mk_disj = foldr1 (HOLogic.mk_binop @{const_name HOL.disj});
-
-fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
-
-
-(* instantiate induction rule *)
-
-fun ind_tac indrule indnames = CSUBGOAL (fn (cgoal, i) =>
- let
- val cert = cterm_of (Thm.theory_of_cterm cgoal);
- val goal = term_of cgoal;
- val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
- val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop (Logic.strip_imp_concl goal));
- 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 (member (op =) indnames o fst);
- fun abstr (t1, t2) =
- (case t1 of
- NONE =>
- (case flt (Term.add_frees t2 []) of
- [(s, T)] => SOME (absfree (s, T) t2)
- | _ => NONE)
- | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))));
- val insts =
- map_filter (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 end);
-
-
-(* perform exhaustive case analysis on last parameter of subgoal i *)
-
-fun exh_tac exh_thm_of = CSUBGOAL (fn (cgoal, i) =>
- let
- val thy = Thm.theory_of_cterm cgoal;
- val goal = term_of cgoal;
- val params = Logic.strip_params goal;
- val (_, Type (tname, _)) = hd (rev params);
- val exhaustion = Thm.lift_rule cgoal (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 (fold_rev (fn (_, T) => fn t => Abs ("z", T, t)) params (Bound 0)))] exhaustion;
- in compose_tac (false, exhaustion', nprems_of exhaustion) i end);
-
-
-(********************** Internal description of datatypes *********************)
-
-datatype dtyp =
- DtTFree of string * sort
- | 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,
- descr : descr,
- inject : thm list,
- distinct : thm list,
- induct : thm,
- inducts : thm list,
- exhaust : thm,
- nchotomy : thm,
- rec_names : string list,
- rec_rewrites : thm list,
- case_name : string,
- case_rewrites : thm list,
- case_cong : thm,
- case_cong_weak : thm,
- split : thm,
- split_asm: thm};
-
-fun mk_Free s T i = Free (s ^ string_of_int i, T);
-
-fun subst_DtTFree _ substs (T as DtTFree a) = the_default T (AList.lookup (op =) substs a)
- | 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 "_"
- (filter_out (equal "") (map name_of_typ Ts) @
- [if Symbol_Pos.is_identifier s' then s' else "x"])
- end
- | name_of_typ _ = "";
-
-fun dtyp_of_typ _ (TFree a) = DtTFree a
- | 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 dest_TFree) Ts = map SOME vs then
- DtRec (find_index (curry op = tname o fst) new_dts)
- else error ("Illegal occurrence of recursive type " ^ quote tname));
-
-fun typ_of_dtyp descr (DtTFree a) = TFree a
- | typ_of_dtyp descr (DtRec i) =
- let val (s, ds, _) = the (AList.lookup (op =) descr i)
- in Type (s, map (typ_of_dtyp descr) ds) end
- | typ_of_dtyp descr (DtType (s, ds)) = Type (s, map (typ_of_dtyp descr) ds);
-
-(* find all non-recursive types in datatype description *)
-
-fun get_nonrec_types descr =
- map (typ_of_dtyp descr) (fold (fn (_, (_, _, constrs)) =>
- fold (fn (_, cargs) => union (op =) (filter_out is_rec_type cargs)) constrs) descr []);
-
-(* get all recursive types in datatype description *)
-
-fun get_rec_types descr = map (fn (_ , (s, ds, _)) =>
- Type (s, map (typ_of_dtyp descr) ds)) descr;
-
-(* get all branching types *)
-
-fun get_branching_types descr =
- map (typ_of_dtyp descr)
- (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) (filter is_rec_type cargs))) constrs)
- descr [];
-
-(* interpret construction of datatype *)
-
-fun interpret_construction descr vs {atyp, dtyp} =
- let
- val typ_of =
- typ_of_dtyp descr #>
- map_atyps (fn TFree (a, _) => TFree (a, the (AList.lookup (op =) vs a)) | T => T);
- fun interpT dT =
- (case strip_dtyp dT of
- (dTs, DtRec l) =>
- let
- val (tyco, dTs', _) = the (AList.lookup (op =) descr l);
- val Ts = map typ_of dTs;
- val Ts' = map typ_of dTs';
- val is_proper = forall (can dest_TFree) Ts';
- in dtyp Ts (l, is_proper) (tyco, Ts') end
- | _ => atyp (typ_of dT));
- fun interpC (c, dTs) = (c, map interpT dTs);
- fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of dTs), map interpC cs);
- in map interpD descr end;
-
-(* nonemptiness check for datatypes *)
-
-fun check_nonempty descr =
- let
- val descr' = flat descr;
- fun is_nonempty_dt is i =
- let
- val (_, _, constrs) = the (AList.lookup (op =) descr' i);
- fun arg_nonempty (_, DtRec i) =
- if member (op =) is i then false
- else is_nonempty_dt (i :: is) i
- | arg_nonempty _ = true;
- in exists (forall (arg_nonempty o strip_dtyp) o snd) constrs end
- val _ = hd descr |> forall (fn (i, (s, _, _)) =>
- is_nonempty_dt [i] i orelse raise Datatype_Empty s)
- in () 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 ctxt orig_descr (dt_info : info Symtab.table) descr i =
- let
- fun typ_error T msg =
- error ("Non-admissible type expression\n" ^
- Syntax.string_of_typ ctxt (typ_of_dtyp (orig_descr @ descr) T) ^ "\n" ^ msg);
-
- fun get_dt_descr T i tname dts =
- (case Symtab.lookup dt_info tname of
- NONE =>
- typ_error T (quote tname ^ " is not a datatype - can't use it in nested recursion")
- | SOME {index, descr, ...} =>
- let
- val (_, vars, _) = the (AList.lookup (op =) descr index);
- val subst = map dest_DtTFree vars ~~ dts
- handle ListPair.UnequalLengths =>
- typ_error T ("Type constructor " ^ quote 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 T (i, Ts, descrs) =
- 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 ctxt orig_descr 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 (cname, cargs) (i, constrs, descrs) =
- let val (i', cargs', descrs') = fold unfold_arg cargs (i, [], descrs)
- in (i', constrs @ [(cname, cargs')], descrs') end;
-
- (* unfold a single datatype *)
-
- fun unfold_datatype (j, (tname, tvars, constrs)) (i, dtypes, descrs) =
- let val (i', constrs', descrs') = fold unfold_constr constrs (i, [], descrs)
- in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs') end;
-
- val (i', descr', descrs) = fold unfold_datatype descr (i, [], []);
-
- in (descr' :: descrs, i') end;
-
-(* find shortest path to constructor with no recursive arguments *)
-
-fun find_nonempty descr is i =
- let
- fun arg_nonempty (_, DtRec i) =
- if member (op =) is i
- then NONE
- else Option.map (Integer.add 1 o snd) (find_nonempty descr (i :: is) i)
- | arg_nonempty _ = SOME 0;
- fun max_inf (SOME i) (SOME j) = SOME (Integer.max i j)
- | max_inf _ _ = NONE;
- fun max xs = fold max_inf xs (SOME 0);
- val (_, _, constrs) = the (AList.lookup (op =) descr i);
- val xs =
- sort (int_ord o pairself snd)
- (map_filter (fn (s, dts) => Option.map (pair s)
- (max (map (arg_nonempty o strip_dtyp) dts))) constrs)
- in if null xs then NONE else SOME (hd xs) end;
-
-fun find_shortest_path descr i = find_nonempty descr [i] i;
-
-end;
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype_codegen.ML
- Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
-
-Code generator facilities for inductive datatypes.
-*)
-
-signature DATATYPE_CODEGEN =
-sig
-end;
-
-structure Datatype_Codegen : DATATYPE_CODEGEN =
-struct
-
-fun add_code_for_datatype fcT_name thy =
- let
- val ctxt = Proof_Context.init_global thy
- val SOME {ctrs, injects, distincts, case_thms, ...} = Ctr_Sugar.ctr_sugar_of ctxt fcT_name
- val Type (_, As) = body_type (fastype_of (hd ctrs))
- in
- Ctr_Sugar_Code.add_ctr_code fcT_name As (map dest_Const ctrs) injects distincts case_thms thy
- end;
-
-val _ = Theory.setup (Datatype_Data.interpretation (K (fold add_code_for_datatype)));
-
-end;
--- a/src/HOL/Tools/Datatype/datatype_data.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,292 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype_data.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package bookkeeping.
-*)
-
-signature DATATYPE_DATA =
-sig
- include DATATYPE_COMMON
-
- val get_all : theory -> info Symtab.table
- val get_info : theory -> string -> info option
- val the_info : theory -> string -> info
- val info_of_constr : theory -> string * typ -> info option
- val info_of_constr_permissive : theory -> string * typ -> info option
- val info_of_case : theory -> string -> info option
- val register: (string * info) list -> theory -> theory
- val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
- val the_descr : theory -> string list ->
- descr * (string * sort) list * string list * string *
- (string list * string list) * (typ list * typ list)
- val all_distincts : theory -> typ list -> thm list list
- val get_constrs : theory -> string -> (string * typ) list option
- val mk_case_names_induct: descr -> attribute
- val mk_case_names_exhausts: descr -> string list -> attribute list
- val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
- val interpretation_data : config * string list -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure Datatype_Data: DATATYPE_DATA =
-struct
-
-(** theory data **)
-
-(* data management *)
-
-structure Data = Theory_Data
-(
- type T =
- {types: Datatype_Aux.info Symtab.table,
- constrs: (string * Datatype_Aux.info) list Symtab.table,
- cases: Datatype_Aux.info Symtab.table};
-
- val empty =
- {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
- val extend = I;
- fun merge
- ({types = types1, constrs = constrs1, cases = cases1},
- {types = types2, constrs = constrs2, cases = cases2}) : T =
- {types = Symtab.merge (K true) (types1, types2),
- constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
- cases = Symtab.merge (K true) (cases1, cases2)};
-);
-
-val get_all = #types o Data.get;
-val get_info = Symtab.lookup o get_all;
-
-fun the_info thy name =
- (case get_info thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
-
-fun info_of_constr thy (c, T) =
- let
- val tab = Symtab.lookup_list (#constrs (Data.get thy)) c;
- in
- (case body_type T of
- Type (tyco, _) => AList.lookup (op =) tab tyco
- | _ => NONE)
- end;
-
-fun info_of_constr_permissive thy (c, T) =
- let
- val tab = Symtab.lookup_list (#constrs (Data.get thy)) c;
- val hint = (case body_type T of Type (tyco, _) => SOME tyco | _ => NONE);
- val default = if null tab then NONE else SOME (snd (List.last tab));
- (*conservative wrt. overloaded constructors*)
- in
- (case hint of
- NONE => default
- | SOME tyco =>
- (case AList.lookup (op =) tab tyco of
- NONE => default (*permissive*)
- | SOME info => SOME info))
- end;
-
-val info_of_case = Symtab.lookup o #cases o Data.get;
-
-fun ctrs_of_exhaust exhaust =
- Logic.strip_imp_prems (prop_of exhaust) |>
- map (head_of o snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o the_single
- o Logic.strip_assums_hyp);
-
-fun case_of_case_rewrite case_rewrite =
- head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of case_rewrite))));
-
-fun ctr_sugar_of_info ({exhaust, nchotomy, inject, distinct, case_rewrites, case_cong,
- case_cong_weak, split, split_asm, ...} : Datatype_Aux.info) =
- {ctrs = ctrs_of_exhaust exhaust,
- casex = case_of_case_rewrite (hd case_rewrites),
- discs = [],
- selss = [],
- exhaust = exhaust,
- nchotomy = nchotomy,
- injects = inject,
- distincts = distinct,
- case_thms = case_rewrites,
- case_cong = case_cong,
- case_cong_weak = case_cong_weak,
- split = split,
- split_asm = split_asm,
- disc_defs = [],
- disc_thmss = [],
- discIs = [],
- sel_defs = [],
- sel_thmss = [],
- distinct_discsss = [],
- exhaust_discs = [],
- exhaust_sels = [],
- collapses = [],
- expands = [],
- split_sels = [],
- split_sel_asms = [],
- case_eq_ifs = []};
-
-fun register dt_infos =
- Data.map (fn {types, constrs, cases} =>
- {types = types |> fold Symtab.update dt_infos,
- constrs = constrs |> fold (fn (constr, dtname_info) =>
- Symtab.map_default (constr, []) (cons dtname_info))
- (maps (fn (dtname, info as {descr, index, ...}) =>
- map (rpair (dtname, info) o fst) (#3 (the (AList.lookup op = descr index)))) dt_infos),
- cases = cases |> fold Symtab.update
- (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)}) #>
- fold (fn (key, info) =>
- Ctr_Sugar.default_register_ctr_sugar_global key (ctr_sugar_of_info info)) dt_infos;
-
-
-(* complex queries *)
-
-fun the_spec thy dtco =
- let
- val {descr, index, ...} = the_info thy dtco;
- val (_, dtys, raw_cos) = the (AList.lookup (op =) descr index);
- val args = map Datatype_Aux.dest_DtTFree dtys;
- val cos = map (fn (co, tys) => (co, map (Datatype_Aux.typ_of_dtyp descr) tys)) raw_cos;
- in (args, cos) end;
-
-fun the_descr thy (raw_tycos as raw_tyco :: _) =
- let
- val info = the_info thy raw_tyco;
- val descr = #descr info;
-
- val (_, dtys, _) = the (AList.lookup (op =) descr (#index info));
- val vs = map Datatype_Aux.dest_DtTFree dtys;
-
- fun is_DtTFree (Datatype_Aux.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 (Datatype_Aux.typ_of_dtyp descr) dTs));
-
- val tycos = map fst dataTs;
- val _ =
- if eq_set (op =) (tycos, raw_tycos) then ()
- else
- error ("Type constructors " ^ commas_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 tycos;
- val (auxnames, _) =
- Name.make_context names
- |> fold_map (Name.variant o Datatype_Aux.name_of_typ) Us;
- val prefix = space_implode "_" names;
-
- in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
-
-fun all_distincts thy Ts =
- let
- fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
- | add_tycos _ = I;
- val tycos = fold add_tycos Ts [];
- in map_filter (Option.map #distinct o get_info thy) tycos end;
-
-fun get_constrs thy dtco =
- (case try (the_spec thy) dtco of
- SOME (args, 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 args);
- fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
- in SOME (map mk_co cos) end
- | NONE => NONE);
-
-
-
-(** various auxiliary **)
-
-(* case names *)
-
-local
-
-fun dt_recs (Datatype_Aux.DtTFree _) = []
- | dt_recs (Datatype_Aux.DtType (_, dts)) = maps dt_recs dts
- | dt_recs (Datatype_Aux.DtRec i) = [i];
-
-fun dt_cases (descr: Datatype_Aux.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 =
- Datatype_Prop.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 = Rule_Cases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (Rule_Cases.case_names o exhaust_cases descr o #1)
- (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
-
-end;
-
-
-
-(** document antiquotation **)
-
-val antiq_setup =
- Thy_Output.antiquotation @{binding datatype} (Args.type_name {proper = true, strict = true})
- (fn {source = src, context = ctxt, ...} => fn dtco =>
- let
- val thy = Proof_Context.theory_of ctxt;
- val (vs, cos) = the_spec thy dtco;
- val ty = Type (dtco, map TFree vs);
- val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt);
- fun pretty_constr (co, tys) =
- Pretty.block (Pretty.breaks
- (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
- map pretty_typ_bracket tys));
- val pretty_datatype =
- Pretty.block
- (Pretty.keyword1 "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
- Thy_Output.output ctxt
- (Thy_Output.maybe_pretty_source (K (K pretty_datatype)) ctxt src [()])
- end);
-
-
-
-(** abstract theory extensions relative to a datatype characterisation **)
-
-structure Datatype_Interpretation = Interpretation
-(
- type T = Datatype_Aux.config * string list;
- val eq: T * T -> bool = eq_snd (op =);
-);
-
-fun with_repaired_path f config (type_names as name :: _) thy =
- thy
- |> Sign.root_path
- |> Sign.add_path (Long_Name.qualifier name)
- |> f config type_names
- |> Sign.restore_naming thy;
-
-fun interpretation f = Datatype_Interpretation.interpretation (uncurry (with_repaired_path f));
-val interpretation_data = Datatype_Interpretation.data;
-
-
-
-(** setup theory **)
-
-val setup =
- antiq_setup #>
- Datatype_Interpretation.init;
-
-open Datatype_Aux;
-
-end;
--- a/src/HOL/Tools/Datatype/datatype_prop.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,427 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype_prop.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package: characteristic properties of datatypes.
-*)
-
-signature DATATYPE_PROP =
-sig
- type descr = Datatype_Aux.descr
- val indexify_names: string list -> string list
- val make_tnames: typ list -> string list
- val make_injs : descr list -> term list list
- val make_distincts : descr list -> term list list (*no symmetric inequalities*)
- val make_ind : descr list -> term
- val make_casedists : descr list -> term list
- val make_primrec_Ts : descr list -> string list -> typ list * typ list
- val make_primrecs : string list -> descr list -> theory -> term list
- val make_cases : string list -> descr list -> theory -> term list list
- val make_splits : string list -> descr list -> theory -> (term * term) list
- val make_case_combs : string list -> descr list -> theory -> string -> term list
- val make_case_cong_weaks : string list -> descr list -> theory -> term list
- val make_case_congs : string list -> descr list -> theory -> term list
- val make_nchotomys : descr list -> term list
-end;
-
-structure Datatype_Prop : DATATYPE_PROP =
-struct
-
-type descr = Datatype_Aux.descr;
-
-
-val indexify_names = Case_Translation.indexify_names;
-val make_tnames = Case_Translation.make_tnames;
-
-fun make_tnames Ts =
- let
- fun type_name (TFree (name, _)) = unprefix "'" name
- | type_name (Type (name, _)) =
- let val name' = Long_Name.base_name name
- in if Symbol_Pos.is_identifier name' then name' else "x" end;
- in indexify_names (map type_name Ts) end;
-
-
-(************************* injectivity of constructors ************************)
-
-fun make_injs descr =
- let
- val descr' = flat descr;
- fun make_inj T (cname, cargs) =
- if null cargs then I
- else
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val constr_t = Const (cname, Ts ---> T);
- val tnames = make_tnames Ts;
- val frees = map Free (tnames ~~ Ts);
- val frees' = map Free (map (suffix "'") 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 @{const_name HOL.conj})
- (map HOLogic.mk_eq (frees ~~ frees')))))
- end;
- in
- map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
- (hd descr) (take (length (hd descr)) (Datatype_Aux.get_rec_types descr'))
- end;
-
-
-(************************* distinctness of constructors ***********************)
-
-fun make_distincts descr =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val newTs = take (length (hd descr)) recTs;
-
- fun prep_constr (cname, cargs) = (cname, map (Datatype_Aux.typ_of_dtyp descr') 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 (suffix "'") (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 =>
- make_distincts' T (map prep_constr constrs)) (hd descr) newTs
- end;
-
-
-(********************************* induction **********************************)
-
-fun make_ind descr =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- 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 (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
- Logic.list_all (map (pair "x") Us,
- HOLogic.mk_Trueprop
- (make_pred (Datatype_Aux.body_index dt) U $
- Datatype_Aux.app_bnds (Free (s, T)) (length Us)))
- end;
-
- val recs = filter Datatype_Aux.is_rec_type cargs;
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val recTs' = map (Datatype_Aux.typ_of_dtyp descr') recs;
- val tnames = Name.variant_list pnames (make_tnames Ts);
- val rec_tnames = map fst (filter (Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs));
- val frees = tnames ~~ Ts;
- val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
- in
- fold_rev (Logic.all o 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 =
- maps (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 @{const_name HOL.conj})
- (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 =
- let
- val descr' = flat descr;
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees;
- in
- fold_rev (Logic.all o 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
- map2 make_casedist (hd descr)
- (take (length (hd descr)) (Datatype_Aux.get_rec_types descr'))
- end;
-
-(*************** characteristic equations for primrec combinator **************)
-
-fun make_primrec_Ts descr used =
- let
- val descr' = flat descr;
-
- val rec_result_Ts =
- map TFree
- (Name.variant_list used (replicate (length descr') "'t") ~~
- replicate (length descr') @{sort type});
-
- val reccomb_fn_Ts = maps (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val recs = filter (Datatype_Aux.is_rec_type o fst) (cargs ~~ Ts);
-
- fun mk_argT (dt, T) =
- binder_types T ---> nth rec_result_Ts (Datatype_Aux.body_index dt);
-
- val argTs = Ts @ map mk_argT recs
- in argTs ---> nth rec_result_Ts i end) constrs) descr';
-
- in (rec_result_Ts, reccomb_fn_Ts) end;
-
-fun make_primrecs reccomb_names descr thy =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val used = fold Term.add_tfree_namesT recTs [];
-
- val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr used;
-
- val rec_fns =
- map (uncurry (Datatype_Aux.mk_Free "f"))
- (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
-
- 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 (cname, cargs) (ts, f :: fs) =
- let
- val recs = filter Datatype_Aux.is_rec_type cargs;
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val recTs' = map (Datatype_Aux.typ_of_dtyp descr') recs;
- val tnames = make_tnames Ts;
- val rec_tnames = map fst (filter (Datatype_Aux.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
- fold_rev (Term.abs o pair "x") Us
- (nth reccombs (Datatype_Aux.body_index dt) $ Datatype_Aux.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
- fold (fn ((dt, T), comb_t) => fold (make_primrec T comb_t) (#3 (snd dt)))
- (descr' ~~ recTs ~~ reccombs) ([], rec_fns)
- |> fst
- end;
-
-(****************** make terms of form t_case f1 ... fn *********************)
-
-fun make_case_combs case_names descr thy fname =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val used = fold Term.add_tfree_namesT recTs [];
- val newTs = take (length (hd descr)) recTs;
- val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
-
- val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs
- in Ts ---> T' end) constrs) (hd descr);
- in
- map (fn ((name, Ts), T) => list_comb
- (Const (name, Ts @ [T] ---> T'),
- map (uncurry (Datatype_Aux.mk_Free fname)) (Ts ~~ (1 upto length Ts))))
- (case_names ~~ case_fn_Ts ~~ newTs)
- end;
-
-(**************** characteristic equations for case combinator ****************)
-
-fun make_cases case_names descr thy =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val newTs = take (length (hd descr)) recTs;
-
- fun make_case T comb_t ((cname, cargs), f) =
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') 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 case_names descr thy "f")
- end;
-
-
-(*************************** the "split" - equations **************************)
-
-fun make_splits case_names descr thy =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val used' = fold Term.add_tfree_namesT recTs [];
- val newTs = take (length (hd descr)) recTs;
- val T' = TFree (singleton (Name.variant_list used') "'t", @{sort type});
- 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 (Datatype_Aux.typ_of_dtyp descr') 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
- (fold_rev (fn Free (s, T) => fn t => HOLogic.mk_all (s, T, t)) frees
- (HOLogic.imp $ eqn $ P') :: t1s,
- fold_rev (fn Free (s, T) => fn t => HOLogic.mk_exists (s, T, t)) frees
- (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) :: t2s)
- end;
-
- val (t1s, t2s) = fold_rev process_constr (constrs ~~ fs) ([], []);
- val lhs = P $ (comb_t $ Free ("x", T));
- in
- (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, Datatype_Aux.mk_conj t1s)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ Datatype_Aux.mk_disj t2s)))
- end
-
- in
- map make_split (hd descr ~~ newTs ~~ make_case_combs case_names descr thy "f")
- end;
-
-(************************* additional rules for TFL ***************************)
-
-fun make_case_cong_weaks case_names descr thy =
- let
- val case_combs = make_case_combs case_names descr 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 case_names descr thy =
- let
- val case_combs = make_case_combs case_names descr thy "f";
- val case_combs' = make_case_combs case_names descr 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 = binder_types (fastype_of f);
- val tnames = Name.variant_list used (make_tnames Ts);
- val frees = map Free (tnames ~~ Ts);
- in
- fold_rev Logic.all frees
- (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 =
- let
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val newTs = take (length (hd descr)) recTs;
-
- fun mk_eqn T (cname, cargs) =
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val tnames = Name.variant_list ["v"] (make_tnames Ts);
- val frees = tnames ~~ Ts;
- in
- fold_rev (fn (s, T') => fn t => HOLogic.mk_exists (s, T', t)) frees
- (HOLogic.mk_eq (Free ("v", T),
- list_comb (Const (cname, Ts ---> T), map Free frees)))
- end;
- in
- map (fn ((_, (_, _, constrs)), T) =>
- HOLogic.mk_Trueprop
- (HOLogic.mk_all ("v", T, Datatype_Aux.mk_disj (map (mk_eqn T) constrs))))
- (hd descr ~~ newTs)
- end;
-
-end;
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,246 +0,0 @@
-(* Title: HOL/Tools/Datatype/datatype_realizer.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Program extraction from proofs involving datatypes:
-realizers for induction and case analysis.
-*)
-
-signature DATATYPE_REALIZER =
-sig
- val add_dt_realizers: Datatype_Aux.config -> string list -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure Datatype_Realizer : DATATYPE_REALIZER =
-struct
-
-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 is_unit t = body_type (fastype_of t) = HOLogic.unitT;
-
-fun tname_of (Type (s, _)) = s
- | tname_of _ = "";
-
-fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Datatype_Aux.info) is thy =
- let
- val ctxt = Proof_Context.init_global thy;
- val cert = cterm_of thy;
-
- val recTs = Datatype_Aux.get_rec_types descr;
- 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 member (op =) is i then TFree ("'" ^ P, @{sort type}) else HOLogic.unitT)
- (descr ~~ pnames);
-
- fun make_pred i T U r x =
- if member (op =) is i then
- Free (nth pnames i, T --> U --> HOLogic.boolT) $ r $ x
- else Free (nth pnames i, U --> HOLogic.boolT) $ x;
-
- fun mk_all i s T t =
- if member (op =) is i then Logic.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 (Datatype_Aux.typ_of_dtyp descr) cargs;
- val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts);
- val recs = filter (Datatype_Aux.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 = Datatype_Aux.mk_Free "f" (map fastype_of vs' ---> rT) j;
- val f' =
- Envir.eta_contract (fold_rev (absfree o dest_Free) vs
- (if member (op =) is i 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 = Datatype_Aux.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
- (Logic.list_all (map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred k rT U (Datatype_Aux.app_bnds r i)
- (Datatype_Aux.app_bnds (Free (s, T)) i))), p)), f)
- end;
- in (apfst (fold_rev (Logic.all o Free) frees) (mk_prems (map Free frees) recs), j + 1) end)
- constrs) (descr ~~ recTs) 1)));
-
- fun mk_proj _ [] 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 = Datatype_Prop.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 (map_filter (fn (((((i, _), T), U), s), tname) =>
- if member (op =) is i 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 @{const_name HOL.conj})
- (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 inst = map (pairself cert) (map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
-
- val thm =
- Goal.prove_internal ctxt (map cert prems) (cert concl)
- (fn prems =>
- EVERY [
- rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]),
- rtac (cterm_instantiate inst induct) 1,
- ALLGOALS (Object_Logic.atomize_prems_tac ctxt),
- rewrite_goals_tac ctxt (@{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)])
- |> Drule.export_without_context;
-
- val ind_name = Thm.derivation_name induct;
- val vs = map (nth pnames) is;
- val (thm', thy') = thy
- |> Sign.root_path
- |> Global_Theory.store_thm
- (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
- ||> Sign.restore_naming thy;
-
- val ivs = rev (Term.add_vars (Logic.varify_global (Datatype_Prop.make_ind [descr])) []);
- val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
- val ivs1 = map Var (filter_out (fn (_, T) => @{type_name bool} = tname_of (body_type T)) ivs);
- val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
-
- val prf =
- Extraction.abs_corr_shyps thy' induct vs ivs2
- (fold_rev (fn (f, p) => fn prf =>
- (case head_of (strip_abs_body f) of
- Free (s, T) =>
- let val T' = Logic.varifyT_global 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)))
- (rec_fns ~~ prems_of thm)
- (Proofterm.proof_combP
- (Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
-
- val r' =
- if null is then r
- else
- Logic.varify_global (fold_rev lambda
- (map Logic.unvarify_global ivs1 @ filter_out is_unit
- (map (head_of o strip_abs_body) rec_fns)) r);
-
- in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
-
-
-fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype_Aux.info)
- thy =
- let
- val ctxt = Proof_Context.init_global thy;
- val cert = cterm_of thy;
- val rT = TFree ("'P", @{sort type});
- val rT' = TVar (("'P", 0), @{sort type});
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr) cargs;
- val frees = Name.variant_list ["P", "y"] (Datatype_Prop.make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees;
- val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
- in
- (r, fold_rev Logic.all free_ts
- (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 = nth (Datatype_Aux.get_rec_types descr) 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.varifyT_global T);
- val y' = Free ("y", T);
-
- val thm =
- Goal.prove_internal ctxt (map cert prems)
- (cert (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y']))))
- (fn prems =>
- EVERY [
- rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1,
- ALLGOALS (EVERY'
- [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites),
- resolve_tac prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
- |> Drule.export_without_context;
-
- val exh_name = Thm.derivation_name exhaust;
- val (thm', thy') = thy
- |> Sign.root_path
- |> Global_Theory.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
- ||> Sign.restore_naming thy;
-
- val P = Var (("P", 0), rT' --> HOLogic.boolT);
- val prf =
- Extraction.abs_corr_shyps thy' exhaust ["P"] [y, P]
- (fold_rev (fn (p, r) => fn prf =>
- Proofterm.forall_intr_proof' (Logic.varify_global r)
- (AbsP ("H", SOME (Logic.varify_global p), prf)))
- (prems ~~ rs)
- (Proofterm.proof_combP
- (Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
- val prf' =
- Extraction.abs_corr_shyps thy' exhaust []
- (map Var (Term.add_vars (prop_of exhaust) [])) (Reconstruct.proof_of exhaust);
- val r' =
- Logic.varify_global (Abs ("y", T,
- (fold_rev (Term.abs o 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'))] thy'
- end;
-
-fun add_dt_realizers config names thy =
- if not (Proofterm.proofs_enabled ()) then thy
- else
- let
- val _ = Datatype_Aux.message config "Adding realizers for induction and case analysis ...";
- val infos = map (Datatype_Data.the_info thy) names;
- val info :: _ = infos;
- in
- thy
- |> fold_rev (make_ind info) (subsets 0 (length (#descr info) - 1))
- |> fold_rev make_casedists infos
- end;
-
-val setup = Datatype_Data.interpretation add_dt_realizers;
-
-end;
--- a/src/HOL/Tools/Datatype/primrec.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,311 +0,0 @@
-(* Title: HOL/Tools/Datatype/primrec.ML
- Author: Norbert Voelker, FernUni Hagen
- Author: Stefan Berghofer, TU Muenchen
- Author: Florian Haftmann, TU Muenchen
-
-Primitive recursive functions on datatypes.
-*)
-
-signature PRIMREC =
-sig
- val add_primrec: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory
- val add_primrec_cmd: (binding * string option * mixfix) list ->
- (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory
- val add_primrec_global: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
- val add_primrec_overloaded: (string * (string * typ) * bool) list ->
- (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
- val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
- local_theory -> (string * (term list * thm list)) * local_theory
-end;
-
-structure Primrec : PRIMREC =
-struct
-
-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);
-
-
-(* preprocessing of equations *)
-
-fun process_eqn is_fixed spec rec_fns =
- let
- val (vs, Ts) = split_list (strip_qnt_vars @{const_name Pure.all} spec);
- val body = strip_qnt_body @{const_name Pure.all} spec;
- val (vs', _) = fold_map Name.variant 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: "
- (Term.add_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 (@{const_name undefined}, dummyT)) :: fns))
- | SOME (ls, cargs', rs, rhs, eq) =>
- let
- val recs = filter (Datatype_Aux.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, (Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
- handle PrimrecError (s, NONE) => primrec_error_eqn s eq
- in
- (fnames'', fnss'', fold_rev absfree (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 (@{const_name undefined},
- replicate (length cargs + length (filter Datatype_Aux.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) (Type.constraint varT raw_rhs);
- in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end;
-
-
-(* find datatypes which contain all datatypes in tnames' *)
-
-fun find_dts _ _ [] = []
- | 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 : Datatype_Aux.info) =>
- if subset (op =) (tnames', 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 ctxt fixes eqs =
- let
- val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed ctxt 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_Data.get_all (Proof_Context.theory_of ctxt)) 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 ctxt fixes fs) raw_defs;
- val names = map snd fnames;
- val names_eqns = map fst eqns;
- val _ =
- if 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 ctxt defs =
- let
- val frees = fold (Variable.add_free_names ctxt) eqs [];
- val rewrites = rec_rewrites' @ map (snd o snd) defs;
- in
- map (fn eq => Goal.prove ctxt frees [] eq
- (fn {context = ctxt', ...} => EVERY [rewrite_goals_tac ctxt' rewrites, rtac refl 1])) 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 ctxt eqn)
- | NONE => ""));
-
-
-(* primrec definition *)
-
-fun add_primrec_simple fixes ts lthy =
- let
- val ((prefix, (_, defs)), prove) = distill lthy fixes ts;
- in
- lthy
- |> fold_map Local_Theory.define defs
- |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs))))
- end;
-
-local
-
-fun gen_primrec 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"), @{attributes [simp, nitpick_simp]});
- in
- lthy
- |> add_primrec_simple fixes (map snd spec)
- |-> (fn (prefix, (ts, simps)) =>
- Spec_Rules.add Spec_Rules.Equational (ts, simps)
- #> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps)
- #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps')
- #>> (fn (_, simps'') => (ts, simps''))))
- end;
-
-in
-
-val add_primrec = gen_primrec Specification.check_spec;
-val add_primrec_cmd = gen_primrec Specification.read_spec;
-
-end;
-
-fun add_primrec_global fixes specs thy =
- let
- val lthy = Named_Target.theory_init thy;
- val ((ts, simps), lthy') = add_primrec fixes specs lthy;
- val simps' = Proof_Context.export lthy' lthy simps;
- in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-fun add_primrec_overloaded ops fixes specs thy =
- let
- val lthy = Overloading.overloading ops thy;
- val ((ts, simps), lthy') = add_primrec fixes specs lthy;
- val simps' = Proof_Context.export lthy' lthy simps;
- in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-end;
--- a/src/HOL/Tools/Datatype/rep_datatype.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,673 +0,0 @@
-(* Title: HOL/Tools/Datatype/rep_datatype.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Representation of existing types as datatypes: proofs and definitions
-independent of concrete representation of datatypes (i.e. requiring
-only abstract properties: injectivity / distinctness of constructors
-and induction).
-*)
-
-signature REP_DATATYPE =
-sig
- val derive_datatype_props : Datatype_Aux.config -> string list -> Datatype_Aux.descr list ->
- thm -> thm list list -> thm list list -> theory -> string list * theory
- val rep_datatype : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
- term list -> theory -> Proof.state
- val rep_datatype_cmd : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
- string list -> theory -> Proof.state
-end;
-
-structure Rep_Datatype: REP_DATATYPE =
-struct
-
-(** derived definitions and proofs **)
-
-(* case distinction theorems *)
-
-fun prove_casedist_thms (config : Datatype_Aux.config)
- new_type_names descr induct case_names_exhausts thy =
- let
- val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
-
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val newTs = take (length (hd descr)) recTs;
-
- val maxidx = Thm.maxidx_of 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 (@{const_name 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 = take i dummyPs @ (P :: drop (i + 1) dummyPs);
- val cert = cterm_of thy;
- val insts' = map cert induct_Ps ~~ map cert insts;
- val induct' =
- refl RS
- (nth (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i RSN (2, rev_mp));
- in
- Goal.prove_sorry_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_index prove_casedist_thm (newTs ~~ Datatype_Prop.make_casedists descr);
- in
- thy
- |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
- (map single case_names_exhausts) casedist_thms
- end;
-
-
-(* primrec combinators *)
-
-fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr
- injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
- let
- val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
-
- val big_name = space_implode "_" new_type_names;
- val thy0 = Sign.add_path big_name thy;
-
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val used = fold Term.add_tfree_namesT recTs [];
- val newTs = take (length (hd descr)) recTs;
-
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
- val big_rec_name' = "rec_set_" ^ big_name;
- val rec_set_names' =
- if length descr' = 1 then [big_rec_name']
- else map (prefix (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) = Datatype_Prop.make_primrec_Ts descr 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 (Datatype_Aux.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 (cname, cargs) (rec_intr_ts, l) =
- let
- fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
- let val free1 = Datatype_Aux.mk_Free "x" U j in
- (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
- ((_, Datatype_Aux.DtRec m), (Us, _)) =>
- let
- val free2 = Datatype_Aux.mk_Free "y" (Us ---> 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, nth rec_sets' m $
- Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
- free1 :: t1s, free2 :: t2s)
- end
- | _ => (j + 1, k, prems, free1 :: t1s, t2s))
- end;
-
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
-
- in
- (rec_intr_ts @
- [Logic.list_implies (prems, HOLogic.mk_Trueprop
- (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
- list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
- end;
-
- val (rec_intr_ts, _) =
- fold (fn ((d, T), set_name) =>
- fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
-
- val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
- thy0
- |> Sign.map_naming Name_Space.conceal
- |> Inductive.add_inductive_global
- {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
- coind = false, no_elim = false, no_ind = true, skip_mono = true}
- (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) []
- ||> Sign.restore_naming thy0;
-
- (* prove uniqueness and termination of primrec combinators *)
-
- val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
-
- fun mk_unique_tac ctxt ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
- let
- val distinct_tac =
- if i < length newTs then
- full_simp_tac (put_simpset HOL_ss ctxt addsimps (nth dist_rewrites i)) 1
- else full_simp_tac (put_simpset HOL_ss ctxt addsimps (flat other_dist_rewrites)) 1;
-
- val inject =
- map (fn r => r RS iffD1)
- (if i < length newTs then nth constr_inject i else injects_of tname);
-
- fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
- let
- val k = length (filter Datatype_Aux.is_rec_type cargs);
- in
- (EVERY
- [DETERM tac,
- REPEAT (etac @{thm ex1E} 1), rtac @{thm 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 ctxt 1,
- REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
- TRY (hyp_subst_tac ctxt 1),
- rtac refl 1,
- REPEAT_DETERM_N (n - j - 1) distinct_tac],
- intrs, j + 1)
- end;
-
- val (tac', intrs', _) =
- fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
- in (tac', intrs') end;
-
- val rec_unique_thms =
- let
- val rec_unique_ts =
- map (fn (((set_t, T1), T2), i) =>
- Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
- absfree ("y", T2) (set_t $ Datatype_Aux.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;
- in
- Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] []
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts))
- (fn {context = ctxt, ...} =>
- #1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
- (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
- rewrite_goals_tac ctxt [mk_meta_eq @{thm choice_eq}], rec_intrs)))))
- end;
-
- val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
-
- (* define primrec combinators *)
-
- val big_reccomb_name = "rec_" ^ space_implode "_" new_type_names;
- val reccomb_names =
- map (Sign.full_bname thy1)
- (if length descr' = 1 then [big_reccomb_name]
- else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto length descr'));
- val reccombs =
- map (fn ((name, T), T') => Const (name, reccomb_fn_Ts @ [T] ---> T'))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- val (reccomb_defs, thy2) =
- thy1
- |> Sign.add_consts (map (fn ((name, T), T') =>
- (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
- (reccomb_names ~~ recTs ~~ rec_result_Ts))
- |> (Global_Theory.add_defs false o map Thm.no_attributes)
- (map
- (fn ((((name, comb), set), T), T') =>
- (Binding.name (Thm.def_name (Long_Name.base_name name)),
- Logic.mk_equals (comb, fold_rev lambda rec_fns (absfree ("x", T)
- (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
- (set $ Free ("x", T) $ Free ("y", T')))))))
- (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
- ||> Sign.parent_path;
-
-
- (* prove characteristic equations for primrec combinators *)
-
- val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
-
- val rec_thms =
- map (fn t =>
- Goal.prove_sorry_global thy2 [] [] t
- (fn {context = ctxt, ...} => EVERY
- [rewrite_goals_tac ctxt reccomb_defs,
- rtac @{thm 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)]))
- (Datatype_Prop.make_primrecs reccomb_names descr thy2);
- in
- thy2
- |> Sign.add_path (space_implode "_" new_type_names)
- |> Global_Theory.note_thmss ""
- [((Binding.name "rec", [Named_Theorems.add @{named_theorems nitpick_simp}]),
- [(rec_thms, [])])]
- ||> Sign.parent_path
- |-> (fn thms => pair (reccomb_names, maps #2 thms))
- end;
-
-
-(* case combinators *)
-
-fun prove_case_thms (config : Datatype_Aux.config)
- new_type_names descr reccomb_names primrec_thms thy =
- let
- val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
-
- val ctxt = Proof_Context.init_global thy;
- val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
-
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val used = fold Term.add_tfree_namesT recTs [];
- val newTs = take (length (hd descr)) recTs;
- val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
-
- fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' dt) ---> T';
-
- val case_dummy_fns =
- map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
- in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
-
- val case_names0 = map (fn s => Sign.full_bname thy1 ("case_" ^ s)) new_type_names;
-
- (* define case combinators via primrec combinators *)
-
- fun def_case ((((i, (_, _, constrs)), T as Type (Tcon, _)), name), recname) (defs, thy) =
- if is_some (Ctr_Sugar.ctr_sugar_of ctxt Tcon) then
- (defs, thy)
- else
- let
- val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
- let
- val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
- val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
- val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
- val frees = take (length cargs) frees';
- val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
- in
- (free, fold_rev (absfree o 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 = flat (take i case_dummy_fns) @ fns2 @ flat (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 (Thm.def_name (Long_Name.base_name name)),
- Logic.mk_equals (Const (name, caseT),
- fold_rev lambda fns1
- (list_comb (reccomb,
- flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
- val ([def_thm], thy') =
- thy
- |> Sign.declare_const_global decl |> snd
- |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
- in (defs @ [def_thm], thy') end;
-
- val (case_defs, thy2) =
- fold def_case (hd descr ~~ newTs ~~ case_names0 ~~ take (length newTs) reccomb_names)
- ([], thy1);
-
- fun prove_case t =
- Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
- EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]);
-
- fun prove_cases (Type (Tcon, _)) ts =
- (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
- SOME {case_thms, ...} => case_thms
- | NONE => map prove_case ts);
-
- val case_thms =
- map2 prove_cases newTs (Datatype_Prop.make_cases case_names0 descr thy2);
-
- fun case_name_of (th :: _) =
- fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
-
- val case_names = map case_name_of case_thms;
- in
- thy2
- |> Context.theory_map
- ((fold o fold) (Named_Theorems.add_thm @{named_theorems nitpick_simp}) case_thms)
- |> Sign.parent_path
- |> Datatype_Aux.store_thmss "case" new_type_names case_thms
- |-> (fn thmss => pair (thmss, case_names))
- end;
-
-
-(* case splitting *)
-
-fun prove_split_thms (config : Datatype_Aux.config)
- new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
- let
- val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
-
- val descr' = flat descr;
- val recTs = Datatype_Aux.get_rec_types descr';
- val newTs = 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;
- fun tac ctxt =
- EVERY [rtac exhaustion' 1,
- ALLGOALS (asm_simp_tac
- (put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))];
- in
- (Goal.prove_sorry_global thy [] [] t1 (tac o #context),
- Goal.prove_sorry_global thy [] [] t2 (tac o #context))
- end;
-
- val split_thm_pairs =
- map prove_split_thms
- (Datatype_Prop.make_splits case_names descr thy ~~ constr_inject ~~
- dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
-
- val (split_thms, split_asm_thms) = split_list split_thm_pairs
-
- in
- thy
- |> Datatype_Aux.store_thms "split" new_type_names split_thms
- ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
- |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
- end;
-
-fun prove_case_cong_weaks new_type_names case_names descr thy =
- let
- fun prove_case_cong_weak t =
- Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
-
- val case_cong_weaks =
- map prove_case_cong_weak (Datatype_Prop.make_case_cong_weaks case_names descr thy);
-
- in thy |> Datatype_Aux.store_thms "case_cong_weak" new_type_names case_cong_weaks end;
-
-
-(* additional theorems for TFL *)
-
-fun prove_nchotomys (config : Datatype_Aux.config) new_type_names descr casedist_thms thy =
- let
- val _ = Datatype_Aux.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 ctxt i 0 =
- EVERY [TRY (rtac disjI1 i), hyp_subst_tac ctxt i, REPEAT (rtac exI i), rtac refl i]
- | tac ctxt i n = rtac disjI2 i THEN tac ctxt i (n - 1);
- in
- Goal.prove_sorry_global thy [] [] t
- (fn {context = ctxt, ...} =>
- EVERY [rtac allI 1,
- Datatype_Aux.exh_tac (K exhaustion) 1,
- ALLGOALS (fn i => tac ctxt i (i - 1))])
- end;
-
- val nchotomys =
- map prove_nchotomy (Datatype_Prop.make_nchotomys descr ~~ casedist_thms);
-
- in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
-
-fun prove_case_congs new_type_names case_names descr nchotomys case_thms thy =
- let
- fun prove_case_cong ((t, nchotomy), case_rewrites) =
- let
- val Const (@{const_name Pure.imp}, _) $ tm $ _ = t;
- val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ 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
- Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {context = ctxt, prems, ...} =>
- let
- val simplify = asm_simp_tac (put_simpset HOL_ss ctxt addsimps (prems @ case_rewrites))
- in
- EVERY [
- simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1,
- cut_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
- (Datatype_Prop.make_case_congs case_names descr thy ~~ nchotomys ~~ case_thms);
-
- in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
-
-
-
-(** derive datatype props **)
-
-local
-
-fun make_dt_info descr induct inducts rec_names rec_rewrites
- (index, (((((((((((_, (tname, _, _))), inject), distinct),
- exhaust), nchotomy), case_name), case_rewrites), case_cong), case_cong_weak),
- (split, split_asm))) =
- (tname,
- {index = index,
- descr = descr,
- inject = inject,
- distinct = distinct,
- induct = induct,
- inducts = inducts,
- exhaust = exhaust,
- nchotomy = nchotomy,
- rec_names = rec_names,
- rec_rewrites = rec_rewrites,
- case_name = case_name,
- case_rewrites = case_rewrites,
- case_cong = case_cong,
- case_cong_weak = case_cong_weak,
- split = split,
- split_asm = split_asm});
-
-in
-
-fun derive_datatype_props config dt_names descr induct inject distinct thy2 =
- let
- val flat_descr = flat descr;
- val new_type_names = map Long_Name.base_name dt_names;
- val _ =
- Datatype_Aux.message config
- ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
-
- val (exhaust, thy3) = thy2
- |> prove_casedist_thms config new_type_names descr induct
- (Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
- val (nchotomys, thy4) = thy3
- |> prove_nchotomys config new_type_names descr exhaust;
- val ((rec_names, rec_rewrites), thy5) = thy4
- |> prove_primrec_thms config new_type_names descr
- (#inject o the o Symtab.lookup (Datatype_Data.get_all thy4)) inject
- (distinct, Datatype_Data.all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr)) induct;
- val ((case_rewrites, case_names), thy6) = thy5
- |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
- val (case_congs, thy7) = thy6
- |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
- val (case_cong_weaks, thy8) = thy7
- |> prove_case_cong_weaks new_type_names case_names descr;
- val (splits, thy9) = thy8
- |> prove_split_thms config new_type_names case_names descr
- inject distinct exhaust case_rewrites;
-
- val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
- val dt_infos =
- map_index
- (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
- (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
- case_names ~~ case_rewrites ~~ case_congs ~~ case_cong_weaks ~~ splits);
- val dt_names = map fst dt_infos;
- val prfx = Binding.qualify true (space_implode "_" new_type_names);
- val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
- val named_rules = flat (map_index (fn (i, tname) =>
- [((Binding.empty, [Induct.induct_type tname]), [([nth inducts i], [])]),
- ((Binding.empty, [Induct.cases_type tname]), [([nth exhaust i], [])])]) dt_names);
- val unnamed_rules = map (fn induct =>
- ((Binding.empty, [Rule_Cases.inner_rule, Induct.induct_type ""]), [([induct], [])]))
- (drop (length dt_names) inducts);
-
- val ctxt = Proof_Context.init_global thy9;
- val case_combs =
- map (Proof_Context.read_const {proper = true, strict = true} ctxt) case_names;
- val constrss = map (fn (dtname, {descr, index, ...}) =>
- map (Proof_Context.read_const {proper = true, strict = true} ctxt o fst)
- (#3 (the (AList.lookup op = descr index)))) dt_infos;
- in
- thy9
- |> Global_Theory.note_thmss ""
- ([((prfx (Binding.name "simps"), []), [(simps, [])]),
- ((prfx (Binding.name "inducts"), []), [(inducts, [])]),
- ((prfx (Binding.name "splits"), []), [(maps (fn (x, y) => [x, y]) splits, [])]),
- ((Binding.empty, [Simplifier.simp_add]),
- [(flat case_rewrites @ flat distinct @ rec_rewrites, [])]),
- ((Binding.empty, [Code.add_default_eqn_attribute]), [(rec_rewrites, [])]),
- ((Binding.empty, [iff_add]), [(flat inject, [])]),
- ((Binding.empty, [Classical.safe_elim NONE]),
- [(map (fn th => th RS notE) (flat distinct), [])]),
- ((Binding.empty, [Simplifier.cong_add]), [(case_cong_weaks, [])]),
- ((Binding.empty, [Induct.induct_simp_add]), [(flat (distinct @ inject), [])])] @
- named_rules @ unnamed_rules)
- |> snd
- |> Datatype_Data.register dt_infos
- |> Context.theory_map (fold2 Case_Translation.register case_combs constrss)
- |> Datatype_Data.interpretation_data (config, dt_names)
- |> pair dt_names
- end;
-
-end;
-
-
-
-(** declare existing type as datatype **)
-
-local
-
-fun prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct thy1 =
- let
- val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
- val new_type_names = map Long_Name.base_name dt_names;
- val prfx = Binding.qualify true (space_implode "_" new_type_names);
- val (((inject, distinct), [(_, [induct])]), thy2) =
- thy1
- |> Datatype_Aux.store_thmss "inject" new_type_names raw_inject
- ||>> Datatype_Aux.store_thmss "distinct" new_type_names raw_distinct
- ||>> Global_Theory.note_thmss ""
- [((prfx (Binding.name "induct"), [Datatype_Data.mk_case_names_induct descr]),
- [([raw_induct], [])])];
- in
- thy2
- |> derive_datatype_props config dt_names [descr] induct inject distinct
- end;
-
-fun gen_rep_datatype prep_term config after_qed raw_ts thy =
- let
- val ctxt = Proof_Context.init_global thy;
-
- fun constr_of_term (Const (c, T)) = (c, T)
- | constr_of_term t = error ("Not a constant: " ^ Syntax.string_of_term ctxt t);
- fun no_constr (c, T) =
- error ("Bad constructor: " ^ Proof_Context.markup_const ctxt c ^ "::" ^
- Syntax.string_of_typ ctxt T);
- fun type_of_constr (cT as (_, T)) =
- let
- val frees = Term.add_tfreesT T [];
- val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_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 (Datatype_Data.get_all thy) tyco then SOME tyco else NONE) raw_cs of
- [] => ()
- | tycos => error ("Type(s) " ^ commas_quote tycos ^ " already represented inductively"));
- 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
- |> foldr1 (Sorts.inter_sort (Sign.classes_of thy));
- val sorts = map inter_sort ms;
- val vs = Name.invent_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 (Datatype_Aux.dtyp_of_typ (map (rpair vs o fst) cs)) o binder_types;
- val dt_names = map fst cs;
-
- fun mk_spec (i, (tyco, constr)) =
- (i, (tyco, map Datatype_Aux.DtTFree vs, (map o apsnd) dtyps_of_typ constr));
- val descr = map_index mk_spec cs;
- val injs = Datatype_Prop.make_injs [descr];
- val half_distincts = Datatype_Prop.make_distincts [descr];
- val ind = Datatype_Prop.make_ind [descr];
- val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
-
- fun after_qed' raw_thms =
- let
- val [[[raw_induct]], raw_inject, half_distinct] =
- unflat rules (map Drule.zero_var_indexes_list raw_thms);
- (*FIXME somehow dubious*)
- in
- Proof_Context.background_theory_result (* FIXME !? *)
- (prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct)
- #-> after_qed
- end;
- in
- ctxt
- |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
- end;
-
-in
-
-val rep_datatype = gen_rep_datatype Sign.cert_term;
-val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global;
-
-end;
-
-
-(* outer syntax *)
-
-val _ =
- Outer_Syntax.command @{command_spec "rep_datatype"} "represent existing types inductively"
- (Scan.repeat1 Parse.term >> (fn ts =>
- Toplevel.theory_to_proof (rep_datatype_cmd Datatype_Aux.default_config (K I) ts)));
-
-end;
--- a/src/HOL/Tools/Function/function.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Function/function.ML Mon Sep 01 16:17:46 2014 +0200
@@ -267,14 +267,14 @@
fun add_case_cong n thy =
let
- val cong = #case_cong (Datatype_Data.the_info thy n)
+ val cong = #case_cong (Old_Datatype_Data.the_info thy n)
|> safe_mk_meta_eq
in
Context.theory_map
(Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
end
-val setup_case_cong = Datatype_Data.interpretation (K (fold add_case_cong))
+val setup_case_cong = Old_Datatype_Data.interpretation (K (fold add_case_cong))
(* setup *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/old_size.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,233 @@
+(* Title: HOL/Tools/Function/old_size.ML
+ Author: Stefan Berghofer, Florian Haftmann, TU Muenchen
+
+Size functions for old-style datatypes.
+*)
+
+signature OLD_SIZE =
+sig
+ val setup: theory -> theory
+end;
+
+structure Old_Size: OLD_SIZE =
+struct
+
+fun plus (t1, t2) = Const (@{const_name Groups.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 _ _ 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 (Old_Datatype_Aux.DtType (name, dts)) =
+ is_some (BNF_LFP_Size.lookup_size_global thy name) andalso exists (is_poly thy) dts
+ | is_poly _ _ = true;
+
+fun constrs_of thy name =
+ let
+ val {descr, index, ...} = Old_Datatype_Data.the_info 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 : Old_Datatype_Aux.info) new_type_names thy =
+ let
+ val {descr, rec_names, rec_rewrites, induct, ...} = info;
+ val l = length new_type_names;
+ val descr' = List.take (descr, l);
+ val tycos = map (#1 o snd) descr';
+ in
+ if forall (fn tyco => can (Sign.arity_sorts thy tyco) [HOLogic.class_size]) tycos then
+ (* nothing to do -- the "size" function is already defined *)
+ thy
+ else
+ let
+ val recTs = Old_Datatype_Aux.get_rec_types descr;
+ val (recTs1, recTs2) = chop l recTs;
+ val (_, (_, paramdts, _)) :: _ = descr;
+ val paramTs = map (Old_Datatype_Aux.typ_of_dtyp descr) paramdts;
+ val ((param_size_fs, param_size_fTs), f_names) = paramTs |>
+ map (fn T as TFree (s, _) =>
+ let
+ val name = "f" ^ unprefix "'" 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 (fst o snd) o BNF_LFP_Size.lookup_size_global thy) |> flat;
+ val extra_size = Option.map fst o BNF_LFP_Size.lookup_size_global thy;
+
+ val (((size_names, size_fns), def_names), def_names') =
+ recTs1 |> map (fn T as Type (s, _) =>
+ let
+ val s' = "size_" ^ Long_Name.base_name s;
+ 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 (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
+ val k = length (filter Old_Datatype_Aux.is_rec_type cargs);
+ val (ts, _, _) = fold_rev (fn ((dt, dt'), T) => fn (us, i, j) =>
+ if Old_Datatype_Aux.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
+ fold_rev (fn T => fn t' => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT) t
+ 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
+ |> Local_Theory.define ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs))
+ |-> (fn (t, (_, thm)) => Spec_Rules.add Spec_Rules.Equational ([t], [thm]) #> pair thm);
+ val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy');
+ val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm;
+ in (thm', lthy') end;
+
+ val ((size_def_thms, size_def_thms'), thy') =
+ thy
+ |> Sign.add_consts (map (fn (s, T) => (Binding.name (Long_Name.base_name s),
+ param_size_fTs @ [T] ---> HOLogic.natT, NoSyn))
+ (size_names ~~ recTs1))
+ |> Global_Theory.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)))
+ ||> Class.instantiation (tycos, 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 []))
+ ||> Local_Theory.exit_global;
+
+ val ctxt = Proof_Context.init_global thy';
+
+ val simpset1 =
+ put_simpset HOL_basic_ss ctxt addsimps @{thm Nat.add_0} :: @{thm Nat.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 (_, 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 Old_Datatype_Aux.split_conj_thm (Goal.prove_sorry ctxt xs []
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj (replicate l @{term True} @
+ map (mk_unfolded_size_eq (AList.lookup op =
+ (new_type_names ~~ map (app fs) rec_combs1)) size_ofp fs)
+ (xs ~~ recTs2 ~~ rec_combs2))))
+ (fn _ => (Old_Datatype_Aux.ind_tac induct 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 (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
+ val tnames = Name.variant_list f_names (Old_Datatype_Prop.make_tnames Ts);
+ val ts = map_filter (fn (sT as (_, 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 =
+ put_simpset HOL_basic_ss ctxt
+ addsimps (rec_rewrites @ size_def_thms @ unfolded_size_eqs1);
+ val simpset3 =
+ put_simpset HOL_basic_ss ctxt
+ 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 => Drule.export_without_context (Goal.prove_sorry 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 Old_Datatype_Aux.is_rec_type overloaded_size_fns (K NONE) simpset3;
+
+ val ([(_, size_thms)], thy'') = thy'
+ |> Global_Theory.note_thmss ""
+ [((Binding.name "size",
+ [Simplifier.simp_add, Named_Theorems.add @{named_theorems nitpick_simp},
+ Thm.declaration_attribute (fn thm =>
+ Context.mapping (Code.add_default_eqn thm) I)]),
+ [(size_eqns, [])])];
+
+ in
+ fold2 (fn new_type_name => fn size_name =>
+ BNF_LFP_Size.register_size_global new_type_name size_name size_thms [])
+ new_type_names size_names thy''
+ end
+ end;
+
+fun add_size_thms _ (new_type_names as name :: _) thy =
+ let
+ val info as {descr, ...} = Old_Datatype_Data.the_info thy name;
+ val prefix = space_implode "_" (map Long_Name.base_name new_type_names);
+ val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt =>
+ Old_Datatype_Aux.is_rec_type dt andalso
+ not (null (fst (Old_Datatype_Aux.strip_dtyp dt)))) cargs) constrs) descr
+ in
+ if no_size then thy
+ else
+ thy
+ |> Sign.add_path prefix
+ |> prove_size_thms info new_type_names
+ |> Sign.restore_naming thy
+ end;
+
+val setup = Old_Datatype_Data.interpretation add_size_thms;
+
+end;
--- a/src/HOL/Tools/Function/size.ML Mon Sep 01 16:17:46 2014 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,233 +0,0 @@
-(* Title: HOL/Tools/Function/size.ML
- Author: Stefan Berghofer, Florian Haftmann, TU Muenchen
-
-Size functions for datatypes.
-*)
-
-signature SIZE =
-sig
- val setup: theory -> theory
-end;
-
-structure Size: SIZE =
-struct
-
-fun plus (t1, t2) = Const (@{const_name Groups.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 _ _ 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 (Datatype_Aux.DtType (name, dts)) =
- is_some (BNF_LFP_Size.lookup_size_global thy name) andalso exists (is_poly thy) dts
- | is_poly _ _ = true;
-
-fun constrs_of thy name =
- let
- val {descr, index, ...} = Datatype_Data.the_info 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_Aux.info) new_type_names thy =
- let
- val {descr, rec_names, rec_rewrites, induct, ...} = info;
- val l = length new_type_names;
- val descr' = List.take (descr, l);
- val tycos = map (#1 o snd) descr';
- in
- if forall (fn tyco => can (Sign.arity_sorts thy tyco) [HOLogic.class_size]) tycos then
- (* nothing to do -- the "size" function is already defined *)
- thy
- else
- let
- val recTs = Datatype_Aux.get_rec_types descr;
- val (recTs1, recTs2) = chop l recTs;
- val (_, (_, paramdts, _)) :: _ = descr;
- val paramTs = map (Datatype_Aux.typ_of_dtyp descr) paramdts;
- val ((param_size_fs, param_size_fTs), f_names) = paramTs |>
- map (fn T as TFree (s, _) =>
- let
- val name = "f" ^ unprefix "'" 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 (fst o snd) o BNF_LFP_Size.lookup_size_global thy) |> flat;
- val extra_size = Option.map fst o BNF_LFP_Size.lookup_size_global thy;
-
- val (((size_names, size_fns), def_names), def_names') =
- recTs1 |> map (fn T as Type (s, _) =>
- let
- val s' = "size_" ^ Long_Name.base_name s;
- 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 (Datatype_Aux.typ_of_dtyp descr) cargs;
- val k = length (filter Datatype_Aux.is_rec_type cargs);
- val (ts, _, _) = fold_rev (fn ((dt, dt'), T) => fn (us, i, j) =>
- if Datatype_Aux.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
- fold_rev (fn T => fn t' => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT) t
- 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
- |> Local_Theory.define ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs))
- |-> (fn (t, (_, thm)) => Spec_Rules.add Spec_Rules.Equational ([t], [thm]) #> pair thm);
- val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy');
- val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm;
- in (thm', lthy') end;
-
- val ((size_def_thms, size_def_thms'), thy') =
- thy
- |> Sign.add_consts (map (fn (s, T) => (Binding.name (Long_Name.base_name s),
- param_size_fTs @ [T] ---> HOLogic.natT, NoSyn))
- (size_names ~~ recTs1))
- |> Global_Theory.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)))
- ||> Class.instantiation (tycos, 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 []))
- ||> Local_Theory.exit_global;
-
- val ctxt = Proof_Context.init_global thy';
-
- val simpset1 =
- put_simpset HOL_basic_ss ctxt addsimps @{thm Nat.add_0} :: @{thm Nat.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 (_, 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 Datatype_Aux.split_conj_thm (Goal.prove_sorry ctxt xs []
- (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj (replicate l @{term True} @
- map (mk_unfolded_size_eq (AList.lookup op =
- (new_type_names ~~ map (app fs) rec_combs1)) size_ofp fs)
- (xs ~~ recTs2 ~~ rec_combs2))))
- (fn _ => (Datatype_Aux.ind_tac induct 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 (Datatype_Aux.typ_of_dtyp descr) cargs;
- val tnames = Name.variant_list f_names (Datatype_Prop.make_tnames Ts);
- val ts = map_filter (fn (sT as (_, 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 =
- put_simpset HOL_basic_ss ctxt
- addsimps (rec_rewrites @ size_def_thms @ unfolded_size_eqs1);
- val simpset3 =
- put_simpset HOL_basic_ss ctxt
- 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 => Drule.export_without_context (Goal.prove_sorry 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 Datatype_Aux.is_rec_type overloaded_size_fns (K NONE) simpset3;
-
- val ([(_, size_thms)], thy'') = thy'
- |> Global_Theory.note_thmss ""
- [((Binding.name "size",
- [Simplifier.simp_add, Named_Theorems.add @{named_theorems nitpick_simp},
- Thm.declaration_attribute (fn thm =>
- Context.mapping (Code.add_default_eqn thm) I)]),
- [(size_eqns, [])])];
-
- in
- fold2 (fn new_type_name => fn size_name =>
- BNF_LFP_Size.register_size_global new_type_name size_name size_thms [])
- new_type_names size_names thy''
- end
- end;
-
-fun add_size_thms _ (new_type_names as name :: _) thy =
- let
- val info as {descr, ...} = Datatype_Data.the_info thy name;
- val prefix = space_implode "_" (map Long_Name.base_name new_type_names);
- val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt =>
- Datatype_Aux.is_rec_type dt andalso
- not (null (fst (Datatype_Aux.strip_dtyp dt)))) cargs) constrs) descr
- in
- if no_size then thy
- else
- thy
- |> Sign.add_path prefix
- |> prove_size_thms info new_type_names
- |> Sign.restore_naming thy
- end;
-
-val setup = Datatype_Data.interpretation add_size_thms;
-
-end;
--- a/src/HOL/Tools/Lifting/lifting_def.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def.ML Mon Sep 01 16:17:46 2014 +0200
@@ -585,7 +585,7 @@
| rename t _ = t
val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt
- val new_names = Datatype_Prop.make_tnames (all_typs fixed_def_t)
+ val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t)
in
rename term new_names
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,802 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package: definitional introduction of datatypes
+with proof of characteristic theorems: injectivity / distinctness
+of constructors and induction. Main interface to datatypes
+after full bootstrap of datatype package.
+*)
+
+signature OLD_DATATYPE =
+sig
+ val distinct_lemma: thm
+ type spec =
+ (binding * (string * sort) list * mixfix) *
+ (binding * typ list * mixfix) list
+ type spec_cmd =
+ (binding * (string * string option) list * mixfix) *
+ (binding * string list * mixfix) list
+ val read_specs: spec_cmd list -> theory -> spec list * Proof.context
+ val check_specs: spec list -> theory -> spec list * Proof.context
+ val add_datatype: Old_Datatype_Aux.config -> spec list -> theory -> string list * theory
+ val add_datatype_cmd: Old_Datatype_Aux.config -> spec_cmd list -> theory -> string list * theory
+ val spec_cmd: spec_cmd parser
+end;
+
+structure Old_Datatype : OLD_DATATYPE =
+struct
+
+(** auxiliary **)
+
+val distinct_lemma = @{lemma "f x \<noteq> f y ==> x \<noteq> y" by iprover};
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+
+fun exh_thm_of (dt_info : Old_Datatype_Aux.info Symtab.table) tname =
+ #exhaust (the (Symtab.lookup dt_info tname));
+
+val In0_inject = @{thm In0_inject};
+val In1_inject = @{thm In1_inject};
+val Scons_inject = @{thm Scons_inject};
+val Leaf_inject = @{thm Leaf_inject};
+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 Lim_inject = @{thm Lim_inject};
+val Inl_inject = @{thm Inl_inject};
+val Inr_inject = @{thm Inr_inject};
+val Suml_inject = @{thm Suml_inject};
+val Sumr_inject = @{thm Sumr_inject};
+
+val datatype_injI =
+ @{lemma "(!!x. ALL y. f x = f y --> x = y) ==> inj f" by (simp add: inj_on_def)};
+
+
+(** proof of characteristic theorems **)
+
+fun representation_proofs (config : Old_Datatype_Aux.config)
+ (dt_info : Old_Datatype_Aux.info Symtab.table) descr types_syntax constr_syntax case_names_induct
+ thy =
+ let
+ val descr' = flat descr;
+ val new_type_names = map (Binding.name_of o fst) types_syntax;
+ val big_name = space_implode "_" new_type_names;
+ val thy1 = Sign.add_path big_name thy;
+ val big_rec_name = "rep_set_" ^ big_name;
+ val rep_set_names' =
+ if length descr' = 1 then [big_rec_name]
+ else map (prefix (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 Old_Datatype_Aux.dest_DtTFree Ts) (hd descr);
+ val leafTs' = Old_Datatype_Aux.get_nonrec_types descr';
+ val branchTs = Old_Datatype_Aux.get_branching_types descr';
+ val branchT =
+ if null branchTs then HOLogic.unitT
+ else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) branchTs;
+ val arities = remove (op =) 0 (Old_Datatype_Aux.get_arities descr');
+ val unneeded_vars =
+ subtract (op =) (fold Term.add_tfreesT (leafTs' @ branchTs) []) (hd tyvars);
+ val leafTs = leafTs' @ map TFree unneeded_vars;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val (newTs, oldTs) = chop (length (hd descr)) recTs;
+ val sumT =
+ if null leafTs then HOLogic.unitT
+ else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) leafTs;
+ val Univ_elT = HOLogic.mk_setT (Type (@{type_name Old_Datatype.node}, [sumT, branchT]));
+ val UnivT = HOLogic.mk_setT Univ_elT;
+ val UnivT' = Univ_elT --> HOLogic.boolT;
+ val Collect = Const (@{const_name Collect}, UnivT' --> UnivT);
+
+ val In0 = Const (@{const_name Old_Datatype.In0}, Univ_elT --> Univ_elT);
+ val In1 = Const (@{const_name Old_Datatype.In1}, Univ_elT --> Univ_elT);
+ val Leaf = Const (@{const_name Old_Datatype.Leaf}, sumT --> Univ_elT);
+ val Lim = Const (@{const_name Old_Datatype.Lim}, (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 (@{const_name Inl}, T1 --> T) $ mk_inj' T1 n2 i
+ else Const (@{const_name Inr}, T2 --> T) $ mk_inj' T2 (n - n2) (i - n2)
+ end;
+ in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) end;
+
+ (* make injections for constructors *)
+
+ fun mk_univ_inj ts = Balanced_Tree.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 @{const_name Old_Datatype.Scons}) 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 (@{const_name Sum_Type.Suml}, mkT T1) $ mk_inj T1 n2 i
+ else Const (@{const_name Sum_Type.Sumr}, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
+ end;
+ in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) end;
+
+ fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t;
+
+ (************** generate introduction rules for representing set **********)
+
+ val _ = Old_Datatype_Aux.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 Old_Datatype_Aux.strip_dtyp dt of
+ (dts, Old_Datatype_Aux.DtRec k) =>
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') dts;
+ val free_t =
+ Old_Datatype_Aux.app_bnds (Old_Datatype_Aux.mk_Free "x" (Ts ---> Univ_elT) j)
+ (length Ts)
+ in
+ (j + 1, Logic.list_all (map (pair "x") Ts,
+ HOLogic.mk_Trueprop
+ (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
+ mk_lim free_t Ts :: ts)
+ end
+ | _ =>
+ let val T = Old_Datatype_Aux.typ_of_dtyp descr' dt
+ in (j + 1, prems, (Leaf $ mk_inj T (Old_Datatype_Aux.mk_Free "x" T j)) :: ts) end);
+
+ val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
+ 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) =
+ thy1
+ |> Sign.map_naming Name_Space.conceal
+ |> Inductive.add_inductive_global
+ {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
+ coind = false, no_elim = true, no_ind = false, skip_mono = true}
+ (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
+ (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
+ ||> Sign.restore_naming thy1;
+
+ (********************************* typedef ********************************)
+
+ val (typedefs, thy3) = thy2
+ |> Sign.parent_path
+ |> fold_map
+ (fn (((name, mx), tvs), c) =>
+ Typedef.add_typedef_global (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 ~~ take (length newTs) rep_set_names)
+ ||> Sign.add_path big_name;
+
+ (*********************** definition of constructors ***********************)
+
+ val big_rep_name = big_name ^ "_Rep_";
+ val rep_names' = map (fn i => big_rep_name ^ string_of_int i) (1 upto length (flat (tl descr)));
+ val all_rep_names =
+ map (#Rep_name o #1 o #2) typedefs @
+ 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 (typedef: Typedef.info) T n
+ ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
+ let
+ fun constr_arg dt (j, l_args, r_args) =
+ let
+ val T = Old_Datatype_Aux.typ_of_dtyp descr' dt;
+ val free_t = Old_Datatype_Aux.mk_Free "x" T j;
+ in
+ (case (Old_Datatype_Aux.strip_dtyp dt, strip_type T) of
+ ((_, Old_Datatype_Aux.DtRec m), (Us, U)) =>
+ (j + 1, free_t :: l_args, mk_lim
+ (Const (nth all_rep_names m, U --> Univ_elT) $
+ Old_Datatype_Aux.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) = fold_rev constr_arg cargs (1, [], []);
+ val constrT = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs ---> T;
+ val ({Abs_name, Rep_name, ...}, _) = typedef;
+ 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 = Thm.def_name (Long_Name.base_name cname);
+ val eqn =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (Rep_name, T --> Univ_elT) $ lhs, rhs));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.add_consts [(cname', constrT, mx)]
+ |> (Global_Theory.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 (((((_, (_, _, constrs)), tname), typedef: Typedef.info), T), constr_syntax)
+ (thy, defs, eqns, rep_congs, dist_lemmas) =
+ let
+ val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
+ val rep_const = cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
+ val cong' = cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong;
+ val dist = cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma;
+ val (thy', defs', eqns', _) =
+ fold (make_constr_def typedef T (length constrs))
+ (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
+ in
+ (Sign.parent_path thy', defs', eqns @ [eqns'],
+ rep_congs @ [cong'], dist_lemmas @ [dist])
+ end;
+
+ val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
+ fold dt_constr_defs
+ (hd descr ~~ new_type_names ~~ map #2 typedefs ~~ newTs ~~ constr_syntax)
+ (thy3 |> Sign.add_consts iso_decls |> Sign.parent_path, [], [], [], []);
+
+
+ (*********** isomorphisms for new types (introduced by typedef) ***********)
+
+ val _ = Old_Datatype_Aux.message config "Proving isomorphism properties ...";
+
+ val collect_simp = rewrite_rule (Proof_Context.init_global thy4) [mk_meta_eq mem_Collect_eq];
+
+ val newT_iso_axms = typedefs |> map (fn (_, (_, {Abs_inverse, Rep_inverse, Rep, ...})) =>
+ (collect_simp Abs_inverse, Rep_inverse, collect_simp Rep));
+
+ val newT_iso_inj_thms = typedefs |> map (fn (_, (_, {Abs_inject, Rep_inject, ...})) =>
+ (collect_simp Abs_inject RS iffD1, Rep_inject RS iffD1));
+
+ (********* isomorphisms between existing types and "unfolded" types *******)
+
+ (*---------------------------------------------------------------------*)
+ (* isomorphisms are defined using primrec-combinators: *)
+ (* generate appropriate functions for instantiating primrec-combinator *)
+ (* *)
+ (* e.g. Rep_dt_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
+ (* *)
+ (* also generate characteristic equations for isomorphisms *)
+ (* *)
+ (* e.g. Rep_dt_i (cons h t) = In1 (Scons (Rep_dt_j h) (Rep_dt_i t)) *)
+ (*---------------------------------------------------------------------*)
+
+ fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
+ let
+ val argTs = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val T = nth recTs k;
+ val rep_const = Const (nth all_rep_names k, T --> Univ_elT);
+ val constr = Const (cname, argTs ---> T);
+
+ fun process_arg ks' dt (i2, i2', ts, Ts) =
+ let
+ val T' = Old_Datatype_Aux.typ_of_dtyp descr' dt;
+ val (Us, U) = strip_type T'
+ in
+ (case Old_Datatype_Aux.strip_dtyp dt of
+ (_, Old_Datatype_Aux.DtRec j) =>
+ if member (op =) ks' j then
+ (i2 + 1, i2' + 1, ts @ [mk_lim (Old_Datatype_Aux.app_bnds
+ (Old_Datatype_Aux.mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
+ Ts @ [Us ---> Univ_elT])
+ else
+ (i2 + 1, i2', ts @ [mk_lim
+ (Const (nth all_rep_names j, U --> Univ_elT) $
+ Old_Datatype_Aux.app_bnds
+ (Old_Datatype_Aux.mk_Free "x" T' i2) (length Us)) Us], Ts)
+ | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (Old_Datatype_Aux.mk_Free "x" T' i2)], Ts))
+ end;
+
+ val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
+ val xs = map (uncurry (Old_Datatype_Aux.mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
+ val ys = map (uncurry (Old_Datatype_Aux.mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
+ val f = fold_rev lambda (xs @ ys) (mk_univ_inj ts n i);
+
+ val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
+ 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 (k, (_, _, constrs)) (fs, eqns, isos) =
+ let
+ val (fs', eqns', _) = fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
+ val iso = (nth recTs k, nth all_rep_names k);
+ in (fs', eqns', isos @ [iso]) end;
+
+ val (fs, eqns, isos) = fold process_dt ds ([], [], []);
+ val fTs = map fastype_of fs;
+ val defs =
+ map (fn (rec_name, (T, iso_name)) =>
+ (Binding.name (Thm.def_name (Long_Name.base_name iso_name)),
+ 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') =
+ (Global_Theory.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 => Goal.prove_sorry_global thy' [] [] eqn
+ (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt rewrites, rtac refl 1])) eqns;
+
+ in (thy', char_thms' @ char_thms) end;
+
+ val (thy5, iso_char_thms) =
+ fold_rev make_iso_defs (tl descr) (Sign.add_path big_name thy4, []);
+
+ (* 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.legacy_freeze prop;
+ val used = Term.add_tfree_names a [];
+
+ fun mk_thm i =
+ let
+ val Ts = map (TFree o rpair @{sort type}) (Name.variant_list used (replicate i "'t"));
+ val f = Free ("f", Ts ---> U);
+ in
+ Goal.prove_sorry_global thy [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.list_all
+ (map (pair "x") Ts, S $ Old_Datatype_Aux.app_bnds f i)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (Term.abs o pair "x") Ts
+ (r $ (a $ Old_Datatype_Aux.app_bnds f i)), f))))
+ (fn _ => EVERY [REPEAT_DETERM_N i (rtac @{thm 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 Rep_dt_i and Rep_dt_i x : rep_set_dt_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 induct = #induct (the (Symtab.lookup dt_info tname));
+
+ fun mk_ind_concl (i, _) =
+ let
+ val T = nth recTs i;
+ val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
+ val rep_set_name = nth rep_set_names i;
+ val concl1 =
+ HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
+ HOLogic.mk_eq (Rep_t $ Old_Datatype_Aux.mk_Free "x" T i, Rep_t $ Bound 0) $
+ HOLogic.mk_eq (Old_Datatype_Aux.mk_Free "x" T i, Bound 0));
+ val concl2 = Const (rep_set_name, UnivT') $ (Rep_t $ Old_Datatype_Aux.mk_Free "x" T i);
+ in (concl1, concl2) end;
+
+ val (ind_concl1, ind_concl2) = split_list (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 =
+ Goal.prove_sorry_global thy5 [] []
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj ind_concl1))
+ (fn {context = ctxt, ...} => EVERY
+ [(Old_Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
+ REPEAT (EVERY
+ [rtac allI 1, rtac impI 1,
+ Old_Datatype_Aux.exh_tac (exh_thm_of dt_info) 1,
+ REPEAT (EVERY
+ [hyp_subst_tac ctxt 1,
+ rewrite_goals_tac ctxt 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 @{thm 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 datatype_injI) (Old_Datatype_Aux.split_conj_thm inj_thm);
+
+ val elem_thm =
+ Goal.prove_sorry_global thy5 [] []
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj ind_concl2))
+ (fn {context = ctxt, ...} =>
+ EVERY [
+ (Old_Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
+ rewrite_goals_tac ctxt 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 @ Old_Datatype_Aux.split_conj_thm elem_thm) end;
+
+ val (iso_inj_thms_unfolded, iso_elem_thms) =
+ fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
+ val iso_inj_thms =
+ map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
+
+ (* prove rep_set_dt_i x --> x : range Rep_dt_i *)
+
+ fun mk_iso_t (((set_name, iso_name), i), T) =
+ let val isoT = T --> Univ_elT in
+ HOLogic.imp $
+ (Const (set_name, UnivT') $ Old_Datatype_Aux.mk_Free "x" Univ_elT i) $
+ (if i < length newTs then @{term True}
+ else HOLogic.mk_mem (Old_Datatype_Aux.mk_Free "x" Univ_elT i,
+ Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
+ Const (iso_name, isoT) $ Const (@{const_abbrev UNIV}, HOLogic.mk_setT T)))
+ end;
+
+ val iso_t = HOLogic.mk_Trueprop (Old_Datatype_Aux.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
+ drop (length newTs) (Old_Datatype_Aux.split_conj_thm
+ (Goal.prove_sorry_global thy5 [] [] iso_t (fn {context = ctxt, ...} => EVERY
+ [(Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW
+ Object_Logic.atomize_prems_tac ctxt) 1,
+ REPEAT (rtac TrueI 1),
+ rewrite_goals_tac ctxt (mk_meta_eq @{thm choice_eq} ::
+ Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
+ rewrite_goals_tac ctxt (map Thm.symmetric range_eqs),
+ REPEAT (EVERY
+ [REPEAT (eresolve_tac ([rangeE, @{thm ex1_implies_ex} RS exE] @
+ maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
+ TRY (hyp_subst_tac ctxt 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 => @{thm f_the_inv_into_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 _ = Old_Datatype_Aux.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
+ Goal.prove_sorry_global thy5 [] [] eqn
+ (fn {context = ctxt, ...} => EVERY
+ [resolve_tac inj_thms 1,
+ rewrite_goals_tac ctxt 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 dist_rewrites' =
+ let
+ fun prove [] = []
+ | prove (t :: ts) =
+ let
+ val dist_thm = Goal.prove_sorry_global thy5 [] [] t (fn {context = ctxt, ...} =>
+ EVERY [simp_tac (put_simpset HOL_ss ctxt addsimps dist_rewrites') 1])
+ in dist_thm :: Drule.zero_var_indexes (dist_thm RS not_sym) :: prove ts end;
+ in prove end;
+
+ val distinct_thms =
+ map2 (prove_distinct_thms) dist_rewrites (Old_Datatype_Prop.make_distincts descr);
+
+ (* 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
+ Goal.prove_sorry_global thy5 [] [] t
+ (fn {context = ctxt, ...} => EVERY
+ [rtac iffI 1,
+ REPEAT (etac conjE 2), hyp_subst_tac ctxt 2, rtac refl 2,
+ dresolve_tac rep_congs 1, dtac @{thm box_equals} 1,
+ REPEAT (resolve_tac rep_thms 1),
+ REPEAT (eresolve_tac inj_thms 1),
+ REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac @{thm 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)
+ (Old_Datatype_Prop.make_injs descr ~~ constr_rep_thms);
+
+ val ((constr_inject', distinct_thms'), thy6) =
+ thy5
+ |> Sign.parent_path
+ |> Old_Datatype_Aux.store_thmss "inject" new_type_names constr_inject
+ ||>> Old_Datatype_Aux.store_thmss "distinct" new_type_names distinct_thms;
+
+ (*************************** induction theorem ****************************)
+
+ val _ = Old_Datatype_Aux.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 @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded;
+ val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
+
+ fun mk_indrule_lemma (i, _) T =
+ let
+ val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ Old_Datatype_Aux.mk_Free "x" T i;
+ val Abs_t =
+ if i < length newTs then
+ Const (#Abs_name (#1 (#2 (nth typedefs i))), Univ_elT --> T)
+ else
+ Const (@{const_name the_inv_into},
+ [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
+ HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT);
+ val prem =
+ HOLogic.imp $
+ (Const (nth rep_set_names i, UnivT') $ Rep_t) $
+ (Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t));
+ val concl =
+ Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $
+ Old_Datatype_Aux.mk_Free "x" T i;
+ in (prem, concl) end;
+
+ val (indrule_lemma_prems, indrule_lemma_concls) =
+ split_list (map2 mk_indrule_lemma descr' recTs);
+
+ val cert = cterm_of thy6;
+
+ val indrule_lemma =
+ Goal.prove_sorry_global thy6 [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (Old_Datatype_Aux.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 = Old_Datatype_Prop.make_ind descr;
+ val dt_induct =
+ Goal.prove_sorry_global thy6 []
+ (Logic.strip_imp_prems dt_induct_prop)
+ (Logic.strip_imp_concl dt_induct_prop)
+ (fn {context = ctxt, prems, ...} =>
+ EVERY
+ [rtac indrule_lemma' 1,
+ (Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW
+ Object_Logic.atomize_prems_tac ctxt) 1,
+ EVERY (map (fn (prem, r) => (EVERY
+ [REPEAT (eresolve_tac Abs_inverse_thms 1),
+ simp_tac (put_simpset HOL_basic_ss ctxt
+ addsimps (Thm.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
+ |> Global_Theory.note_thmss ""
+ [((Binding.qualify true big_name (Binding.name "induct"), [case_names_induct]),
+ [([dt_induct], [])])];
+ in
+ ((constr_inject', distinct_thms', dt_induct'), thy7)
+ end;
+
+
+
+(** datatype definition **)
+
+(* specifications *)
+
+type spec = (binding * (string * sort) list * mixfix) * (binding * typ list * mixfix) list;
+
+type spec_cmd =
+ (binding * (string * string option) list * mixfix) * (binding * string list * mixfix) list;
+
+local
+
+fun parse_spec ctxt ((b, args, mx), constrs) =
+ ((b, map (apsnd (Typedecl.read_constraint ctxt)) args, mx),
+ constrs |> map (fn (c, Ts, mx') => (c, map (Syntax.parse_typ ctxt) Ts, mx')));
+
+fun check_specs ctxt (specs: spec list) =
+ let
+ fun prep_spec ((tname, args, mx), constrs) tys =
+ let
+ val (args', tys1) = chop (length args) tys;
+ val (constrs', tys3) = (constrs, tys1) |-> fold_map (fn (cname, cargs, mx') => fn tys2 =>
+ let val (cargs', tys3) = chop (length cargs) tys2;
+ in ((cname, cargs', mx'), tys3) end);
+ in (((tname, map dest_TFree args', mx), constrs'), tys3) end;
+
+ val all_tys =
+ specs |> maps (fn ((_, args, _), cs) => map TFree args @ maps #2 cs)
+ |> Syntax.check_typs ctxt;
+
+ in #1 (fold_map prep_spec specs all_tys) end;
+
+fun prep_specs parse raw_specs thy =
+ let
+ val ctxt = thy
+ |> Sign.add_types_global (map (fn ((b, args, mx), _) => (b, length args, mx)) raw_specs)
+ |> Proof_Context.init_global
+ |> fold (fn ((_, args, _), _) => fold (fn (a, _) =>
+ Variable.declare_typ (TFree (a, dummyS))) args) raw_specs;
+ val specs = check_specs ctxt (map (parse ctxt) raw_specs);
+ in (specs, ctxt) end;
+
+in
+
+val read_specs = prep_specs parse_spec;
+val check_specs = prep_specs (K I);
+
+end;
+
+
+(* main commands *)
+
+fun gen_add_datatype prep_specs config raw_specs thy =
+ let
+ val _ = Theory.requires thy (Context.theory_name @{theory}) "datatype definitions";
+
+ val (dts, spec_ctxt) = prep_specs raw_specs thy;
+ val ((_, tyvars, _), _) :: _ = dts;
+ val string_of_tyvar = Syntax.string_of_typ spec_ctxt o TFree;
+
+ val (new_dts, types_syntax) = dts |> map (fn ((tname, tvs, mx), _) =>
+ let val full_tname = Sign.full_name thy tname in
+ (case duplicates (op =) tvs of
+ [] =>
+ if eq_set (op =) (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 " ^ Binding.print tname ^
+ " : " ^ commas (map string_of_tyvar dups)))
+ end) |> split_list;
+ val dt_names = map fst new_dts;
+
+ val _ =
+ (case duplicates (op =) (map fst new_dts) of
+ [] => ()
+ | dups => error ("Duplicate datatypes: " ^ commas_quote dups));
+
+ fun prep_dt_spec ((tname, tvs, _), constrs) (dts', constr_syntax, i) =
+ let
+ fun prep_constr (cname, cargs, mx) (constrs, constr_syntax') =
+ let
+ val _ =
+ (case subtract (op =) tvs (fold Term.add_tfreesT cargs []) of
+ [] => ()
+ | vs => error ("Extra type variables on rhs: " ^ commas (map string_of_tyvar vs)));
+ val c = Sign.full_name_path thy (Binding.name_of tname) cname;
+ in
+ (constrs @ [(c, map (Old_Datatype_Aux.dtyp_of_typ new_dts) cargs)],
+ constr_syntax' @ [(cname, mx)])
+ end handle ERROR msg =>
+ cat_error msg ("The error above occurred in constructor " ^ Binding.print cname ^
+ " of datatype " ^ Binding.print tname);
+
+ val (constrs', constr_syntax') = fold prep_constr constrs ([], []);
+ in
+ (case duplicates (op =) (map fst constrs') of
+ [] =>
+ (dts' @ [(i, (Sign.full_name thy tname, map Old_Datatype_Aux.DtTFree tvs, constrs'))],
+ constr_syntax @ [constr_syntax'], i + 1)
+ | dups =>
+ error ("Duplicate constructors " ^ commas_quote dups ^
+ " in datatype " ^ Binding.print tname))
+ end;
+
+ val (dts', constr_syntax, i) = fold prep_dt_spec dts ([], [], 0);
+
+ val dt_info = Old_Datatype_Data.get_all thy;
+ val (descr, _) = Old_Datatype_Aux.unfold_datatypes spec_ctxt dts' dt_info dts' i;
+ val _ =
+ Old_Datatype_Aux.check_nonempty descr
+ handle (exn as Old_Datatype_Aux.Datatype_Empty s) =>
+ if #strict config then error ("Nonemptiness check failed for datatype " ^ quote s)
+ else reraise exn;
+
+ val _ =
+ Old_Datatype_Aux.message config
+ ("Constructing datatype(s) " ^ commas_quote (map (Binding.name_of o #1 o #1) dts));
+ in
+ thy
+ |> representation_proofs config dt_info descr types_syntax constr_syntax
+ (Old_Datatype_Data.mk_case_names_induct (flat descr))
+ |-> (fn (inject, distinct, induct) =>
+ Old_Rep_Datatype.derive_datatype_props config dt_names descr induct inject distinct)
+ end;
+
+val add_datatype = gen_add_datatype check_specs;
+val add_datatype_cmd = gen_add_datatype read_specs;
+
+
+(* outer syntax *)
+
+val spec_cmd =
+ Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
+ (@{keyword "="} |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix))
+ >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Parse.triple1 cons));
+
+val _ =
+ Outer_Syntax.command @{command_spec "datatype"} "define inductive datatypes"
+ (Parse.and_list1 spec_cmd
+ >> (Toplevel.theory o (snd oo add_datatype_cmd Old_Datatype_Aux.default_config)));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,402 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype_aux.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package: auxiliary data structures and functions.
+*)
+
+signature OLD_DATATYPE_COMMON =
+sig
+ type config = {strict : bool, quiet : bool}
+ val default_config : config
+ datatype dtyp =
+ DtTFree of string * sort
+ | DtType of string * dtyp list
+ | DtRec of int
+ type descr = (int * (string * dtyp list * (string * dtyp list) list)) list
+ type info =
+ {index : int,
+ descr : descr,
+ inject : thm list,
+ distinct : thm list,
+ induct : thm,
+ inducts : thm list,
+ exhaust : thm,
+ nchotomy : thm,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ case_name : string,
+ case_rewrites : thm list,
+ case_cong : thm,
+ case_cong_weak : thm,
+ split : thm,
+ split_asm: thm}
+end
+
+signature OLD_DATATYPE_AUX =
+sig
+ include OLD_DATATYPE_COMMON
+
+ val message : config -> string -> unit
+
+ 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 ind_tac : thm -> string list -> int -> tactic
+ val exh_tac : (string -> thm) -> int -> tactic
+
+ exception Datatype
+ exception Datatype_Empty of string
+ val name_of_typ : typ -> string
+ val dtyp_of_typ : (string * (string * sort) list) list -> typ -> dtyp
+ val mk_Free : string -> typ -> int -> term
+ val is_rec_type : dtyp -> bool
+ val typ_of_dtyp : descr -> dtyp -> typ
+ val dest_DtTFree : dtyp -> string * sort
+ 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 -> typ list
+ val get_branching_types : descr -> typ list
+ val get_arities : descr -> int list
+ val get_rec_types : descr -> typ list
+ val interpret_construction : descr -> (string * sort) list ->
+ {atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a} ->
+ ((string * typ list) * (string * 'a list) list) list
+ val check_nonempty : descr list -> unit
+ val unfold_datatypes : Proof.context -> descr -> info Symtab.table ->
+ descr -> int -> descr list * int
+ val find_shortest_path : descr -> int -> (string * int) option
+end;
+
+structure Old_Datatype_Aux : OLD_DATATYPE_AUX =
+struct
+
+(* datatype option flags *)
+
+type config = {strict : bool, quiet : bool};
+val default_config : config = {strict = true, quiet = false};
+
+fun message ({quiet = true, ...} : config) s = writeln s
+ | message _ _ = ();
+
+
+(* store theorems in theory *)
+
+fun store_thmss_atts name tnames attss thmss =
+ fold_map (fn ((tname, atts), thms) =>
+ Global_Theory.note_thmss ""
+ [((Binding.qualify true tname (Binding.name name), atts), [(thms, [])])]
+ #-> (fn [(_, res)] => pair res)) (tnames ~~ attss ~~ thmss);
+
+fun store_thmss name tnames = store_thmss_atts name tnames (replicate (length tnames) []);
+
+fun store_thms_atts name tnames attss thms =
+ fold_map (fn ((tname, atts), thm) =>
+ Global_Theory.note_thmss ""
+ [((Binding.qualify true tname (Binding.name name), atts), [([thm], [])])]
+ #-> (fn [(_, [res])] => pair res)) (tnames ~~ attss ~~ thms);
+
+fun store_thms name tnames = store_thms_atts name 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 @{const_name HOL.conj});
+val mk_disj = foldr1 (HOLogic.mk_binop @{const_name HOL.disj});
+
+fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
+
+
+(* instantiate induction rule *)
+
+fun ind_tac indrule indnames = CSUBGOAL (fn (cgoal, i) =>
+ let
+ val cert = cterm_of (Thm.theory_of_cterm cgoal);
+ val goal = term_of cgoal;
+ val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
+ val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop (Logic.strip_imp_concl goal));
+ 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 (member (op =) indnames o fst);
+ fun abstr (t1, t2) =
+ (case t1 of
+ NONE =>
+ (case flt (Term.add_frees t2 []) of
+ [(s, T)] => SOME (absfree (s, T) t2)
+ | _ => NONE)
+ | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))));
+ val insts =
+ map_filter (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 end);
+
+
+(* perform exhaustive case analysis on last parameter of subgoal i *)
+
+fun exh_tac exh_thm_of = CSUBGOAL (fn (cgoal, i) =>
+ let
+ val thy = Thm.theory_of_cterm cgoal;
+ val goal = term_of cgoal;
+ val params = Logic.strip_params goal;
+ val (_, Type (tname, _)) = hd (rev params);
+ val exhaustion = Thm.lift_rule cgoal (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 (fold_rev (fn (_, T) => fn t => Abs ("z", T, t)) params (Bound 0)))] exhaustion;
+ in compose_tac (false, exhaustion', nprems_of exhaustion) i end);
+
+
+(********************** Internal description of datatypes *********************)
+
+datatype dtyp =
+ DtTFree of string * sort
+ | 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,
+ descr : descr,
+ inject : thm list,
+ distinct : thm list,
+ induct : thm,
+ inducts : thm list,
+ exhaust : thm,
+ nchotomy : thm,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ case_name : string,
+ case_rewrites : thm list,
+ case_cong : thm,
+ case_cong_weak : thm,
+ split : thm,
+ split_asm: thm};
+
+fun mk_Free s T i = Free (s ^ string_of_int i, T);
+
+fun subst_DtTFree _ substs (T as DtTFree a) = the_default T (AList.lookup (op =) substs a)
+ | 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 "_"
+ (filter_out (equal "") (map name_of_typ Ts) @
+ [if Symbol_Pos.is_identifier s' then s' else "x"])
+ end
+ | name_of_typ _ = "";
+
+fun dtyp_of_typ _ (TFree a) = DtTFree a
+ | 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 dest_TFree) Ts = map SOME vs then
+ DtRec (find_index (curry op = tname o fst) new_dts)
+ else error ("Illegal occurrence of recursive type " ^ quote tname));
+
+fun typ_of_dtyp descr (DtTFree a) = TFree a
+ | typ_of_dtyp descr (DtRec i) =
+ let val (s, ds, _) = the (AList.lookup (op =) descr i)
+ in Type (s, map (typ_of_dtyp descr) ds) end
+ | typ_of_dtyp descr (DtType (s, ds)) = Type (s, map (typ_of_dtyp descr) ds);
+
+(* find all non-recursive types in datatype description *)
+
+fun get_nonrec_types descr =
+ map (typ_of_dtyp descr) (fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => union (op =) (filter_out is_rec_type cargs)) constrs) descr []);
+
+(* get all recursive types in datatype description *)
+
+fun get_rec_types descr = map (fn (_ , (s, ds, _)) =>
+ Type (s, map (typ_of_dtyp descr) ds)) descr;
+
+(* get all branching types *)
+
+fun get_branching_types descr =
+ map (typ_of_dtyp descr)
+ (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) (filter is_rec_type cargs))) constrs)
+ descr [];
+
+(* interpret construction of datatype *)
+
+fun interpret_construction descr vs {atyp, dtyp} =
+ let
+ val typ_of =
+ typ_of_dtyp descr #>
+ map_atyps (fn TFree (a, _) => TFree (a, the (AList.lookup (op =) vs a)) | T => T);
+ fun interpT dT =
+ (case strip_dtyp dT of
+ (dTs, DtRec l) =>
+ let
+ val (tyco, dTs', _) = the (AList.lookup (op =) descr l);
+ val Ts = map typ_of dTs;
+ val Ts' = map typ_of dTs';
+ val is_proper = forall (can dest_TFree) Ts';
+ in dtyp Ts (l, is_proper) (tyco, Ts') end
+ | _ => atyp (typ_of dT));
+ fun interpC (c, dTs) = (c, map interpT dTs);
+ fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of dTs), map interpC cs);
+ in map interpD descr end;
+
+(* nonemptiness check for datatypes *)
+
+fun check_nonempty descr =
+ let
+ val descr' = flat descr;
+ fun is_nonempty_dt is i =
+ let
+ val (_, _, constrs) = the (AList.lookup (op =) descr' i);
+ fun arg_nonempty (_, DtRec i) =
+ if member (op =) is i then false
+ else is_nonempty_dt (i :: is) i
+ | arg_nonempty _ = true;
+ in exists (forall (arg_nonempty o strip_dtyp) o snd) constrs end
+ val _ = hd descr |> forall (fn (i, (s, _, _)) =>
+ is_nonempty_dt [i] i orelse raise Datatype_Empty s)
+ in () 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 ctxt orig_descr (dt_info : info Symtab.table) descr i =
+ let
+ fun typ_error T msg =
+ error ("Non-admissible type expression\n" ^
+ Syntax.string_of_typ ctxt (typ_of_dtyp (orig_descr @ descr) T) ^ "\n" ^ msg);
+
+ fun get_dt_descr T i tname dts =
+ (case Symtab.lookup dt_info tname of
+ NONE =>
+ typ_error T (quote tname ^ " is not a datatype - can't use it in nested recursion")
+ | SOME {index, descr, ...} =>
+ let
+ val (_, vars, _) = the (AList.lookup (op =) descr index);
+ val subst = map dest_DtTFree vars ~~ dts
+ handle ListPair.UnequalLengths =>
+ typ_error T ("Type constructor " ^ quote 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 T (i, Ts, descrs) =
+ 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 ctxt orig_descr 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 (cname, cargs) (i, constrs, descrs) =
+ let val (i', cargs', descrs') = fold unfold_arg cargs (i, [], descrs)
+ in (i', constrs @ [(cname, cargs')], descrs') end;
+
+ (* unfold a single datatype *)
+
+ fun unfold_datatype (j, (tname, tvars, constrs)) (i, dtypes, descrs) =
+ let val (i', constrs', descrs') = fold unfold_constr constrs (i, [], descrs)
+ in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs') end;
+
+ val (i', descr', descrs) = fold unfold_datatype descr (i, [], []);
+
+ in (descr' :: descrs, i') end;
+
+(* find shortest path to constructor with no recursive arguments *)
+
+fun find_nonempty descr is i =
+ let
+ fun arg_nonempty (_, DtRec i) =
+ if member (op =) is i
+ then NONE
+ else Option.map (Integer.add 1 o snd) (find_nonempty descr (i :: is) i)
+ | arg_nonempty _ = SOME 0;
+ fun max_inf (SOME i) (SOME j) = SOME (Integer.max i j)
+ | max_inf _ _ = NONE;
+ fun max xs = fold max_inf xs (SOME 0);
+ val (_, _, constrs) = the (AList.lookup (op =) descr i);
+ val xs =
+ sort (int_ord o pairself snd)
+ (map_filter (fn (s, dts) => Option.map (pair s)
+ (max (map (arg_nonempty o strip_dtyp) dts))) constrs)
+ in if null xs then NONE else SOME (hd xs) end;
+
+fun find_shortest_path descr i = find_nonempty descr [i] i;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_codegen.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,25 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype_codegen.ML
+ Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
+
+Code generator facilities for inductive datatypes.
+*)
+
+signature OLD_DATATYPE_CODEGEN =
+sig
+end;
+
+structure Old_Datatype_Codegen : OLD_DATATYPE_CODEGEN =
+struct
+
+fun add_code_for_datatype fcT_name thy =
+ let
+ val ctxt = Proof_Context.init_global thy
+ val SOME {ctrs, injects, distincts, case_thms, ...} = Ctr_Sugar.ctr_sugar_of ctxt fcT_name
+ val Type (_, As) = body_type (fastype_of (hd ctrs))
+ in
+ Ctr_Sugar_Code.add_ctr_code fcT_name As (map dest_Const ctrs) injects distincts case_thms thy
+ end;
+
+val _ = Theory.setup (Old_Datatype_Data.interpretation (K (fold add_code_for_datatype)));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_data.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,292 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype_data.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package bookkeeping.
+*)
+
+signature OLD_DATATYPE_DATA =
+sig
+ include OLD_DATATYPE_COMMON
+
+ val get_all : theory -> info Symtab.table
+ val get_info : theory -> string -> info option
+ val the_info : theory -> string -> info
+ val info_of_constr : theory -> string * typ -> info option
+ val info_of_constr_permissive : theory -> string * typ -> info option
+ val info_of_case : theory -> string -> info option
+ val register: (string * info) list -> theory -> theory
+ val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
+ val the_descr : theory -> string list ->
+ descr * (string * sort) list * string list * string *
+ (string list * string list) * (typ list * typ list)
+ val all_distincts : theory -> typ list -> thm list list
+ val get_constrs : theory -> string -> (string * typ) list option
+ val mk_case_names_induct: descr -> attribute
+ val mk_case_names_exhausts: descr -> string list -> attribute list
+ val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
+ val interpretation_data : config * string list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure Old_Datatype_Data: OLD_DATATYPE_DATA =
+struct
+
+(** theory data **)
+
+(* data management *)
+
+structure Data = Theory_Data
+(
+ type T =
+ {types: Old_Datatype_Aux.info Symtab.table,
+ constrs: (string * Old_Datatype_Aux.info) list Symtab.table,
+ cases: Old_Datatype_Aux.info Symtab.table};
+
+ val empty =
+ {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
+ val extend = I;
+ fun merge
+ ({types = types1, constrs = constrs1, cases = cases1},
+ {types = types2, constrs = constrs2, cases = cases2}) : T =
+ {types = Symtab.merge (K true) (types1, types2),
+ constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
+ cases = Symtab.merge (K true) (cases1, cases2)};
+);
+
+val get_all = #types o Data.get;
+val get_info = Symtab.lookup o get_all;
+
+fun the_info thy name =
+ (case get_info thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
+
+fun info_of_constr thy (c, T) =
+ let
+ val tab = Symtab.lookup_list (#constrs (Data.get thy)) c;
+ in
+ (case body_type T of
+ Type (tyco, _) => AList.lookup (op =) tab tyco
+ | _ => NONE)
+ end;
+
+fun info_of_constr_permissive thy (c, T) =
+ let
+ val tab = Symtab.lookup_list (#constrs (Data.get thy)) c;
+ val hint = (case body_type T of Type (tyco, _) => SOME tyco | _ => NONE);
+ val default = if null tab then NONE else SOME (snd (List.last tab));
+ (*conservative wrt. overloaded constructors*)
+ in
+ (case hint of
+ NONE => default
+ | SOME tyco =>
+ (case AList.lookup (op =) tab tyco of
+ NONE => default (*permissive*)
+ | SOME info => SOME info))
+ end;
+
+val info_of_case = Symtab.lookup o #cases o Data.get;
+
+fun ctrs_of_exhaust exhaust =
+ Logic.strip_imp_prems (prop_of exhaust) |>
+ map (head_of o snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o the_single
+ o Logic.strip_assums_hyp);
+
+fun case_of_case_rewrite case_rewrite =
+ head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of case_rewrite))));
+
+fun ctr_sugar_of_info ({exhaust, nchotomy, inject, distinct, case_rewrites, case_cong,
+ case_cong_weak, split, split_asm, ...} : Old_Datatype_Aux.info) =
+ {ctrs = ctrs_of_exhaust exhaust,
+ casex = case_of_case_rewrite (hd case_rewrites),
+ discs = [],
+ selss = [],
+ exhaust = exhaust,
+ nchotomy = nchotomy,
+ injects = inject,
+ distincts = distinct,
+ case_thms = case_rewrites,
+ case_cong = case_cong,
+ case_cong_weak = case_cong_weak,
+ split = split,
+ split_asm = split_asm,
+ disc_defs = [],
+ disc_thmss = [],
+ discIs = [],
+ sel_defs = [],
+ sel_thmss = [],
+ distinct_discsss = [],
+ exhaust_discs = [],
+ exhaust_sels = [],
+ collapses = [],
+ expands = [],
+ split_sels = [],
+ split_sel_asms = [],
+ case_eq_ifs = []};
+
+fun register dt_infos =
+ Data.map (fn {types, constrs, cases} =>
+ {types = types |> fold Symtab.update dt_infos,
+ constrs = constrs |> fold (fn (constr, dtname_info) =>
+ Symtab.map_default (constr, []) (cons dtname_info))
+ (maps (fn (dtname, info as {descr, index, ...}) =>
+ map (rpair (dtname, info) o fst) (#3 (the (AList.lookup op = descr index)))) dt_infos),
+ cases = cases |> fold Symtab.update
+ (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)}) #>
+ fold (fn (key, info) =>
+ Ctr_Sugar.default_register_ctr_sugar_global key (ctr_sugar_of_info info)) dt_infos;
+
+
+(* complex queries *)
+
+fun the_spec thy dtco =
+ let
+ val {descr, index, ...} = the_info thy dtco;
+ val (_, dtys, raw_cos) = the (AList.lookup (op =) descr index);
+ val args = map Old_Datatype_Aux.dest_DtTFree dtys;
+ val cos = map (fn (co, tys) => (co, map (Old_Datatype_Aux.typ_of_dtyp descr) tys)) raw_cos;
+ in (args, cos) end;
+
+fun the_descr thy (raw_tycos as raw_tyco :: _) =
+ let
+ val info = the_info thy raw_tyco;
+ val descr = #descr info;
+
+ val (_, dtys, _) = the (AList.lookup (op =) descr (#index info));
+ val vs = map Old_Datatype_Aux.dest_DtTFree dtys;
+
+ fun is_DtTFree (Old_Datatype_Aux.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 (Old_Datatype_Aux.typ_of_dtyp descr) dTs));
+
+ val tycos = map fst dataTs;
+ val _ =
+ if eq_set (op =) (tycos, raw_tycos) then ()
+ else
+ error ("Type constructors " ^ commas_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 tycos;
+ val (auxnames, _) =
+ Name.make_context names
+ |> fold_map (Name.variant o Old_Datatype_Aux.name_of_typ) Us;
+ val prefix = space_implode "_" names;
+
+ in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
+
+fun all_distincts thy Ts =
+ let
+ fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
+ | add_tycos _ = I;
+ val tycos = fold add_tycos Ts [];
+ in map_filter (Option.map #distinct o get_info thy) tycos end;
+
+fun get_constrs thy dtco =
+ (case try (the_spec thy) dtco of
+ SOME (args, 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 args);
+ fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
+ in SOME (map mk_co cos) end
+ | NONE => NONE);
+
+
+
+(** various auxiliary **)
+
+(* case names *)
+
+local
+
+fun dt_recs (Old_Datatype_Aux.DtTFree _) = []
+ | dt_recs (Old_Datatype_Aux.DtType (_, dts)) = maps dt_recs dts
+ | dt_recs (Old_Datatype_Aux.DtRec i) = [i];
+
+fun dt_cases (descr: Old_Datatype_Aux.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 =
+ Old_Datatype_Prop.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 = Rule_Cases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (Rule_Cases.case_names o exhaust_cases descr o #1)
+ (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
+
+end;
+
+
+
+(** document antiquotation **)
+
+val antiq_setup =
+ Thy_Output.antiquotation @{binding datatype} (Args.type_name {proper = true, strict = true})
+ (fn {source = src, context = ctxt, ...} => fn dtco =>
+ let
+ val thy = Proof_Context.theory_of ctxt;
+ val (vs, cos) = the_spec thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt);
+ fun pretty_constr (co, tys) =
+ Pretty.block (Pretty.breaks
+ (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
+ map pretty_typ_bracket tys));
+ val pretty_datatype =
+ Pretty.block
+ (Pretty.keyword1 "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
+ Thy_Output.output ctxt
+ (Thy_Output.maybe_pretty_source (K (K pretty_datatype)) ctxt src [()])
+ end);
+
+
+
+(** abstract theory extensions relative to a datatype characterisation **)
+
+structure Datatype_Interpretation = Interpretation
+(
+ type T = Old_Datatype_Aux.config * string list;
+ val eq: T * T -> bool = eq_snd (op =);
+);
+
+fun with_repaired_path f config (type_names as name :: _) thy =
+ thy
+ |> Sign.root_path
+ |> Sign.add_path (Long_Name.qualifier name)
+ |> f config type_names
+ |> Sign.restore_naming thy;
+
+fun interpretation f = Datatype_Interpretation.interpretation (uncurry (with_repaired_path f));
+val interpretation_data = Datatype_Interpretation.data;
+
+
+
+(** setup theory **)
+
+val setup =
+ antiq_setup #>
+ Datatype_Interpretation.init;
+
+open Old_Datatype_Aux;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_prop.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,428 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype_prop.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package: characteristic properties of datatypes.
+*)
+
+signature OLD_DATATYPE_PROP =
+sig
+ type descr = Old_Datatype_Aux.descr
+ val indexify_names: string list -> string list
+ val make_tnames: typ list -> string list
+ val make_injs : descr list -> term list list
+ val make_distincts : descr list -> term list list (*no symmetric inequalities*)
+ val make_ind : descr list -> term
+ val make_casedists : descr list -> term list
+ val make_primrec_Ts : descr list -> string list -> typ list * typ list
+ val make_primrecs : string list -> descr list -> theory -> term list
+ val make_cases : string list -> descr list -> theory -> term list list
+ val make_splits : string list -> descr list -> theory -> (term * term) list
+ val make_case_combs : string list -> descr list -> theory -> string -> term list
+ val make_case_cong_weaks : string list -> descr list -> theory -> term list
+ val make_case_congs : string list -> descr list -> theory -> term list
+ val make_nchotomys : descr list -> term list
+end;
+
+structure Old_Datatype_Prop : OLD_DATATYPE_PROP =
+struct
+
+type descr = Old_Datatype_Aux.descr;
+
+
+val indexify_names = Case_Translation.indexify_names;
+val make_tnames = Case_Translation.make_tnames;
+
+fun make_tnames Ts =
+ let
+ fun type_name (TFree (name, _)) = unprefix "'" name
+ | type_name (Type (name, _)) =
+ let val name' = Long_Name.base_name name
+ in if Symbol_Pos.is_identifier name' then name' else "x" end;
+ in indexify_names (map type_name Ts) end;
+
+
+(************************* injectivity of constructors ************************)
+
+fun make_injs descr =
+ let
+ val descr' = flat descr;
+ fun make_inj T (cname, cargs) =
+ if null cargs then I
+ else
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val constr_t = Const (cname, Ts ---> T);
+ val tnames = make_tnames Ts;
+ val frees = map Free (tnames ~~ Ts);
+ val frees' = map Free (map (suffix "'") 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 @{const_name HOL.conj})
+ (map HOLogic.mk_eq (frees ~~ frees')))))
+ end;
+ in
+ map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
+ (hd descr) (take (length (hd descr)) (Old_Datatype_Aux.get_rec_types descr'))
+ end;
+
+
+(************************* distinctness of constructors ***********************)
+
+fun make_distincts descr =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val newTs = take (length (hd descr)) recTs;
+
+ fun prep_constr (cname, cargs) = (cname, map (Old_Datatype_Aux.typ_of_dtyp descr') 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 (suffix "'") (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 =>
+ make_distincts' T (map prep_constr constrs)) (hd descr) newTs
+ end;
+
+
+(********************************* induction **********************************)
+
+fun make_ind descr =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ 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 (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
+ Logic.list_all (map (pair "x") Us,
+ HOLogic.mk_Trueprop
+ (make_pred (Old_Datatype_Aux.body_index dt) U $
+ Old_Datatype_Aux.app_bnds (Free (s, T)) (length Us)))
+ end;
+
+ val recs = filter Old_Datatype_Aux.is_rec_type cargs;
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val recTs' = map (Old_Datatype_Aux.typ_of_dtyp descr') recs;
+ val tnames = Name.variant_list pnames (make_tnames Ts);
+ val rec_tnames = map fst (filter (Old_Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs));
+ val frees = tnames ~~ Ts;
+ val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
+ in
+ fold_rev (Logic.all o 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 =
+ maps (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 @{const_name HOL.conj})
+ (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 =
+ let
+ val descr' = flat descr;
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees;
+ in
+ fold_rev (Logic.all o 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
+ map2 make_casedist (hd descr)
+ (take (length (hd descr)) (Old_Datatype_Aux.get_rec_types descr'))
+ end;
+
+(*************** characteristic equations for primrec combinator **************)
+
+fun make_primrec_Ts descr used =
+ let
+ val descr' = flat descr;
+
+ val rec_result_Ts =
+ map TFree
+ (Name.variant_list used (replicate (length descr') "'t") ~~
+ replicate (length descr') @{sort type});
+
+ val reccomb_fn_Ts = maps (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val recs = filter (Old_Datatype_Aux.is_rec_type o fst) (cargs ~~ Ts);
+
+ fun mk_argT (dt, T) =
+ binder_types T ---> nth rec_result_Ts (Old_Datatype_Aux.body_index dt);
+
+ val argTs = Ts @ map mk_argT recs
+ in argTs ---> nth rec_result_Ts i end) constrs) descr';
+
+ in (rec_result_Ts, reccomb_fn_Ts) end;
+
+fun make_primrecs reccomb_names descr thy =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val used = fold Term.add_tfree_namesT recTs [];
+
+ val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr used;
+
+ val rec_fns =
+ map (uncurry (Old_Datatype_Aux.mk_Free "f"))
+ (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
+
+ 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 (cname, cargs) (ts, f :: fs) =
+ let
+ val recs = filter Old_Datatype_Aux.is_rec_type cargs;
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val recTs' = map (Old_Datatype_Aux.typ_of_dtyp descr') recs;
+ val tnames = make_tnames Ts;
+ val rec_tnames = map fst (filter (Old_Datatype_Aux.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
+ fold_rev (Term.abs o pair "x") Us
+ (nth reccombs (Old_Datatype_Aux.body_index dt) $
+ Old_Datatype_Aux.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
+ fold (fn ((dt, T), comb_t) => fold (make_primrec T comb_t) (#3 (snd dt)))
+ (descr' ~~ recTs ~~ reccombs) ([], rec_fns)
+ |> fst
+ end;
+
+(****************** make terms of form t_case f1 ... fn *********************)
+
+fun make_case_combs case_names descr thy fname =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val used = fold Term.add_tfree_namesT recTs [];
+ val newTs = take (length (hd descr)) recTs;
+ val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
+
+ val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs
+ in Ts ---> T' end) constrs) (hd descr);
+ in
+ map (fn ((name, Ts), T) => list_comb
+ (Const (name, Ts @ [T] ---> T'),
+ map (uncurry (Old_Datatype_Aux.mk_Free fname)) (Ts ~~ (1 upto length Ts))))
+ (case_names ~~ case_fn_Ts ~~ newTs)
+ end;
+
+(**************** characteristic equations for case combinator ****************)
+
+fun make_cases case_names descr thy =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val newTs = take (length (hd descr)) recTs;
+
+ fun make_case T comb_t ((cname, cargs), f) =
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') 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 case_names descr thy "f")
+ end;
+
+
+(*************************** the "split" - equations **************************)
+
+fun make_splits case_names descr thy =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val used' = fold Term.add_tfree_namesT recTs [];
+ val newTs = take (length (hd descr)) recTs;
+ val T' = TFree (singleton (Name.variant_list used') "'t", @{sort type});
+ 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 (Old_Datatype_Aux.typ_of_dtyp descr') 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
+ (fold_rev (fn Free (s, T) => fn t => HOLogic.mk_all (s, T, t)) frees
+ (HOLogic.imp $ eqn $ P') :: t1s,
+ fold_rev (fn Free (s, T) => fn t => HOLogic.mk_exists (s, T, t)) frees
+ (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) :: t2s)
+ end;
+
+ val (t1s, t2s) = fold_rev process_constr (constrs ~~ fs) ([], []);
+ val lhs = P $ (comb_t $ Free ("x", T));
+ in
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, Old_Datatype_Aux.mk_conj t1s)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ Old_Datatype_Aux.mk_disj t2s)))
+ end
+
+ in
+ map make_split (hd descr ~~ newTs ~~ make_case_combs case_names descr thy "f")
+ end;
+
+(************************* additional rules for TFL ***************************)
+
+fun make_case_cong_weaks case_names descr thy =
+ let
+ val case_combs = make_case_combs case_names descr 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 case_names descr thy =
+ let
+ val case_combs = make_case_combs case_names descr thy "f";
+ val case_combs' = make_case_combs case_names descr 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 = binder_types (fastype_of f);
+ val tnames = Name.variant_list used (make_tnames Ts);
+ val frees = map Free (tnames ~~ Ts);
+ in
+ fold_rev Logic.all frees
+ (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 =
+ let
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val newTs = take (length (hd descr)) recTs;
+
+ fun mk_eqn T (cname, cargs) =
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val tnames = Name.variant_list ["v"] (make_tnames Ts);
+ val frees = tnames ~~ Ts;
+ in
+ fold_rev (fn (s, T') => fn t => HOLogic.mk_exists (s, T', t)) frees
+ (HOLogic.mk_eq (Free ("v", T),
+ list_comb (Const (cname, Ts ---> T), map Free frees)))
+ end;
+ in
+ map (fn ((_, (_, _, constrs)), T) =>
+ HOLogic.mk_Trueprop
+ (HOLogic.mk_all ("v", T, Old_Datatype_Aux.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/Old_Datatype/old_datatype_realizer.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,246 @@
+(* Title: HOL/Tools/Old_Datatype/old_datatype_realizer.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Program extraction from proofs involving datatypes:
+realizers for induction and case analysis.
+*)
+
+signature OLD_DATATYPE_REALIZER =
+sig
+ val add_dt_realizers: Old_Datatype_Aux.config -> string list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure Old_Datatype_Realizer : OLD_DATATYPE_REALIZER =
+struct
+
+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 is_unit t = body_type (fastype_of t) = HOLogic.unitT;
+
+fun tname_of (Type (s, _)) = s
+ | tname_of _ = "";
+
+fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Old_Datatype_Aux.info) is thy =
+ let
+ val ctxt = Proof_Context.init_global thy;
+ val cert = cterm_of thy;
+
+ val recTs = Old_Datatype_Aux.get_rec_types descr;
+ 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 member (op =) is i then TFree ("'" ^ P, @{sort type}) else HOLogic.unitT)
+ (descr ~~ pnames);
+
+ fun make_pred i T U r x =
+ if member (op =) is i then
+ Free (nth pnames i, T --> U --> HOLogic.boolT) $ r $ x
+ else Free (nth pnames i, U --> HOLogic.boolT) $ x;
+
+ fun mk_all i s T t =
+ if member (op =) is i then Logic.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 (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
+ val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts);
+ val recs = filter (Old_Datatype_Aux.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 = Old_Datatype_Aux.mk_Free "f" (map fastype_of vs' ---> rT) j;
+ val f' =
+ Envir.eta_contract (fold_rev (absfree o dest_Free) vs
+ (if member (op =) is i 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 = Old_Datatype_Aux.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
+ (Logic.list_all (map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred k rT U (Old_Datatype_Aux.app_bnds r i)
+ (Old_Datatype_Aux.app_bnds (Free (s, T)) i))), p)), f)
+ end;
+ in (apfst (fold_rev (Logic.all o Free) frees) (mk_prems (map Free frees) recs), j + 1) end)
+ constrs) (descr ~~ recTs) 1)));
+
+ fun mk_proj _ [] 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 = Old_Datatype_Prop.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 (map_filter (fn (((((i, _), T), U), s), tname) =>
+ if member (op =) is i 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 @{const_name HOL.conj})
+ (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 inst = map (pairself cert) (map head_of (HOLogic.dest_conj
+ (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
+
+ val thm =
+ Goal.prove_internal ctxt (map cert prems) (cert concl)
+ (fn prems =>
+ EVERY [
+ rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]),
+ rtac (cterm_instantiate inst induct) 1,
+ ALLGOALS (Object_Logic.atomize_prems_tac ctxt),
+ rewrite_goals_tac ctxt (@{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)])
+ |> Drule.export_without_context;
+
+ val ind_name = Thm.derivation_name induct;
+ val vs = map (nth pnames) is;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> Global_Theory.store_thm
+ (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
+ ||> Sign.restore_naming thy;
+
+ val ivs = rev (Term.add_vars (Logic.varify_global (Old_Datatype_Prop.make_ind [descr])) []);
+ val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
+ val ivs1 = map Var (filter_out (fn (_, T) => @{type_name bool} = tname_of (body_type T)) ivs);
+ val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
+
+ val prf =
+ Extraction.abs_corr_shyps thy' induct vs ivs2
+ (fold_rev (fn (f, p) => fn prf =>
+ (case head_of (strip_abs_body f) of
+ Free (s, T) =>
+ let val T' = Logic.varifyT_global 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)))
+ (rec_fns ~~ prems_of thm)
+ (Proofterm.proof_combP
+ (Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
+
+ val r' =
+ if null is then r
+ else
+ Logic.varify_global (fold_rev lambda
+ (map Logic.unvarify_global ivs1 @ filter_out is_unit
+ (map (head_of o strip_abs_body) rec_fns)) r);
+
+ in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
+
+
+fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Old_Datatype_Aux.info)
+ thy =
+ let
+ val ctxt = Proof_Context.init_global thy;
+ val cert = cterm_of thy;
+ val rT = TFree ("'P", @{sort type});
+ val rT' = TVar (("'P", 0), @{sort type});
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
+ val frees = Name.variant_list ["P", "y"] (Old_Datatype_Prop.make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees;
+ val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
+ in
+ (r, fold_rev Logic.all free_ts
+ (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 = nth (Old_Datatype_Aux.get_rec_types descr) 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.varifyT_global T);
+ val y' = Free ("y", T);
+
+ val thm =
+ Goal.prove_internal ctxt (map cert prems)
+ (cert (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y']))))
+ (fn prems =>
+ EVERY [
+ rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1,
+ ALLGOALS (EVERY'
+ [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites),
+ resolve_tac prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
+ |> Drule.export_without_context;
+
+ val exh_name = Thm.derivation_name exhaust;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> Global_Theory.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
+ ||> Sign.restore_naming thy;
+
+ val P = Var (("P", 0), rT' --> HOLogic.boolT);
+ val prf =
+ Extraction.abs_corr_shyps thy' exhaust ["P"] [y, P]
+ (fold_rev (fn (p, r) => fn prf =>
+ Proofterm.forall_intr_proof' (Logic.varify_global r)
+ (AbsP ("H", SOME (Logic.varify_global p), prf)))
+ (prems ~~ rs)
+ (Proofterm.proof_combP
+ (Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
+ val prf' =
+ Extraction.abs_corr_shyps thy' exhaust []
+ (map Var (Term.add_vars (prop_of exhaust) [])) (Reconstruct.proof_of exhaust);
+ val r' =
+ Logic.varify_global (Abs ("y", T,
+ (fold_rev (Term.abs o 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'))] thy'
+ end;
+
+fun add_dt_realizers config names thy =
+ if not (Proofterm.proofs_enabled ()) then thy
+ else
+ let
+ val _ = Old_Datatype_Aux.message config "Adding realizers for induction and case analysis ...";
+ val infos = map (Old_Datatype_Data.the_info thy) names;
+ val info :: _ = infos;
+ in
+ thy
+ |> fold_rev (make_ind info) (subsets 0 (length (#descr info) - 1))
+ |> fold_rev make_casedists infos
+ end;
+
+val setup = Old_Datatype_Data.interpretation add_dt_realizers;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_primrec.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,311 @@
+(* Title: HOL/Tools/Old_Datatype/old_primrec.ML
+ Author: Norbert Voelker, FernUni Hagen
+ Author: Stefan Berghofer, TU Muenchen
+ Author: Florian Haftmann, TU Muenchen
+
+Primitive recursive functions on datatypes.
+*)
+
+signature OLD_PRIMREC =
+sig
+ val add_primrec: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory
+ val add_primrec_cmd: (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory
+ val add_primrec_global: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
+ val add_primrec_overloaded: (string * (string * typ) * bool) list ->
+ (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
+ val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+ local_theory -> (string * (term list * thm list)) * local_theory
+end;
+
+structure Old_Primrec : OLD_PRIMREC =
+struct
+
+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);
+
+
+(* preprocessing of equations *)
+
+fun process_eqn is_fixed spec rec_fns =
+ let
+ val (vs, Ts) = split_list (strip_qnt_vars @{const_name Pure.all} spec);
+ val body = strip_qnt_body @{const_name Pure.all} spec;
+ val (vs', _) = fold_map Name.variant 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: "
+ (Term.add_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 (@{const_name undefined}, dummyT)) :: fns))
+ | SOME (ls, cargs', rs, rhs, eq) =>
+ let
+ val recs = filter (Old_Datatype_Aux.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, (Old_Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
+ handle PrimrecError (s, NONE) => primrec_error_eqn s eq
+ in
+ (fnames'', fnss'', fold_rev absfree (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 (@{const_name undefined},
+ replicate (length cargs + length (filter Old_Datatype_Aux.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) (Type.constraint varT raw_rhs);
+ in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end;
+
+
+(* find datatypes which contain all datatypes in tnames' *)
+
+fun find_dts _ _ [] = []
+ | 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 : Old_Datatype_Aux.info) =>
+ if subset (op =) (tnames', 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 ctxt fixes eqs =
+ let
+ val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed ctxt 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 (Old_Datatype_Data.get_all (Proof_Context.theory_of ctxt)) 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 ctxt fixes fs) raw_defs;
+ val names = map snd fnames;
+ val names_eqns = map fst eqns;
+ val _ =
+ if 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 ctxt defs =
+ let
+ val frees = fold (Variable.add_free_names ctxt) eqs [];
+ val rewrites = rec_rewrites' @ map (snd o snd) defs;
+ in
+ map (fn eq => Goal.prove ctxt frees [] eq
+ (fn {context = ctxt', ...} => EVERY [rewrite_goals_tac ctxt' rewrites, rtac refl 1])) 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 ctxt eqn)
+ | NONE => ""));
+
+
+(* primrec definition *)
+
+fun add_primrec_simple fixes ts lthy =
+ let
+ val ((prefix, (_, defs)), prove) = distill lthy fixes ts;
+ in
+ lthy
+ |> fold_map Local_Theory.define defs
+ |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs))))
+ end;
+
+local
+
+fun gen_primrec 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"), @{attributes [simp, nitpick_simp]});
+ in
+ lthy
+ |> add_primrec_simple fixes (map snd spec)
+ |-> (fn (prefix, (ts, simps)) =>
+ Spec_Rules.add Spec_Rules.Equational (ts, simps)
+ #> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps)
+ #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps')
+ #>> (fn (_, simps'') => (ts, simps''))))
+ end;
+
+in
+
+val add_primrec = gen_primrec Specification.check_spec;
+val add_primrec_cmd = gen_primrec Specification.read_spec;
+
+end;
+
+fun add_primrec_global fixes specs thy =
+ let
+ val lthy = Named_Target.theory_init thy;
+ val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+ val simps' = Proof_Context.export lthy' lthy simps;
+ in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+fun add_primrec_overloaded ops fixes specs thy =
+ let
+ val lthy = Overloading.overloading ops thy;
+ val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+ val simps' = Proof_Context.export lthy' lthy simps;
+ in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Mon Sep 01 16:17:46 2014 +0200
@@ -0,0 +1,680 @@
+(* Title: HOL/Tools/Old_Datatype/old_rep_datatype.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Representation of existing types as datatypes: proofs and definitions
+independent of concrete representation of datatypes (i.e. requiring
+only abstract properties: injectivity / distinctness of constructors
+and induction).
+*)
+
+signature OLD_REP_DATATYPE =
+sig
+ val derive_datatype_props : Old_Datatype_Aux.config -> string list ->
+ Old_Datatype_Aux.descr list -> thm -> thm list list -> thm list list -> theory ->
+ string list * theory
+ val rep_datatype : Old_Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
+ term list -> theory -> Proof.state
+ val rep_datatype_cmd : Old_Datatype_Aux.config ->
+ (string list -> Proof.context -> Proof.context) -> string list -> theory -> Proof.state
+end;
+
+structure Old_Rep_Datatype: OLD_REP_DATATYPE =
+struct
+
+(** derived definitions and proofs **)
+
+(* case distinction theorems *)
+
+fun prove_casedist_thms (config : Old_Datatype_Aux.config)
+ new_type_names descr induct case_names_exhausts thy =
+ let
+ val _ = Old_Datatype_Aux.message config "Proving case distinction theorems ...";
+
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val newTs = take (length (hd descr)) recTs;
+
+ val maxidx = Thm.maxidx_of 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 (@{const_name 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 = take i dummyPs @ (P :: drop (i + 1) dummyPs);
+ val cert = cterm_of thy;
+ val insts' = map cert induct_Ps ~~ map cert insts;
+ val induct' =
+ refl RS
+ (nth (Old_Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i
+ RSN (2, rev_mp));
+ in
+ Goal.prove_sorry_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_index prove_casedist_thm (newTs ~~ Old_Datatype_Prop.make_casedists descr);
+ in
+ thy
+ |> Old_Datatype_Aux.store_thms_atts "exhaust" new_type_names
+ (map single case_names_exhausts) casedist_thms
+ end;
+
+
+(* primrec combinators *)
+
+fun prove_primrec_thms (config : Old_Datatype_Aux.config) new_type_names descr
+ injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
+ let
+ val _ = Old_Datatype_Aux.message config "Constructing primrec combinators ...";
+
+ val big_name = space_implode "_" new_type_names;
+ val thy0 = Sign.add_path big_name thy;
+
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val used = fold Term.add_tfree_namesT recTs [];
+ val newTs = take (length (hd descr)) recTs;
+
+ val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+ val big_rec_name' = "rec_set_" ^ big_name;
+ val rec_set_names' =
+ if length descr' = 1 then [big_rec_name']
+ else map (prefix (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) = Old_Datatype_Prop.make_primrec_Ts descr 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 (Old_Datatype_Aux.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 (cname, cargs) (rec_intr_ts, l) =
+ let
+ fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
+ let val free1 = Old_Datatype_Aux.mk_Free "x" U j in
+ (case (Old_Datatype_Aux.strip_dtyp dt, strip_type U) of
+ ((_, Old_Datatype_Aux.DtRec m), (Us, _)) =>
+ let
+ val free2 = Old_Datatype_Aux.mk_Free "y" (Us ---> 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, nth rec_sets' m $
+ Old_Datatype_Aux.app_bnds free1 i $
+ Old_Datatype_Aux.app_bnds free2 i)) :: prems,
+ free1 :: t1s, free2 :: t2s)
+ end
+ | _ => (j + 1, k, prems, free1 :: t1s, t2s))
+ end;
+
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
+
+ in
+ (rec_intr_ts @
+ [Logic.list_implies (prems, HOLogic.mk_Trueprop
+ (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
+ list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
+ end;
+
+ val (rec_intr_ts, _) =
+ fold (fn ((d, T), set_name) =>
+ fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
+
+ val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
+ thy0
+ |> Sign.map_naming Name_Space.conceal
+ |> Inductive.add_inductive_global
+ {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
+ coind = false, no_elim = false, no_ind = true, skip_mono = true}
+ (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) []
+ ||> Sign.restore_naming thy0;
+
+ (* prove uniqueness and termination of primrec combinators *)
+
+ val _ = Old_Datatype_Aux.message config
+ "Proving termination and uniqueness of primrec functions ...";
+
+ fun mk_unique_tac ctxt ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
+ let
+ val distinct_tac =
+ if i < length newTs then
+ full_simp_tac (put_simpset HOL_ss ctxt addsimps (nth dist_rewrites i)) 1
+ else full_simp_tac (put_simpset HOL_ss ctxt addsimps (flat other_dist_rewrites)) 1;
+
+ val inject =
+ map (fn r => r RS iffD1)
+ (if i < length newTs then nth constr_inject i else injects_of tname);
+
+ fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
+ let
+ val k = length (filter Old_Datatype_Aux.is_rec_type cargs);
+ in
+ (EVERY
+ [DETERM tac,
+ REPEAT (etac @{thm ex1E} 1), rtac @{thm 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 ctxt 1,
+ REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
+ TRY (hyp_subst_tac ctxt 1),
+ rtac refl 1,
+ REPEAT_DETERM_N (n - j - 1) distinct_tac],
+ intrs, j + 1)
+ end;
+
+ val (tac', intrs', _) =
+ fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
+ in (tac', intrs') end;
+
+ val rec_unique_thms =
+ let
+ val rec_unique_ts =
+ map (fn (((set_t, T1), T2), i) =>
+ Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
+ absfree ("y", T2) (set_t $ Old_Datatype_Aux.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;
+ in
+ Old_Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] []
+ (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj rec_unique_ts))
+ (fn {context = ctxt, ...} =>
+ #1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
+ (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
+ rewrite_goals_tac ctxt [mk_meta_eq @{thm choice_eq}], rec_intrs)))))
+ end;
+
+ val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
+
+ (* define primrec combinators *)
+
+ val big_reccomb_name = "rec_" ^ space_implode "_" new_type_names;
+ val reccomb_names =
+ map (Sign.full_bname thy1)
+ (if length descr' = 1 then [big_reccomb_name]
+ else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto length descr'));
+ val reccombs =
+ map (fn ((name, T), T') => Const (name, reccomb_fn_Ts @ [T] ---> T'))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ val (reccomb_defs, thy2) =
+ thy1
+ |> Sign.add_consts (map (fn ((name, T), T') =>
+ (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts))
+ |> (Global_Theory.add_defs false o map Thm.no_attributes)
+ (map
+ (fn ((((name, comb), set), T), T') =>
+ (Binding.name (Thm.def_name (Long_Name.base_name name)),
+ Logic.mk_equals (comb, fold_rev lambda rec_fns (absfree ("x", T)
+ (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
+ (set $ Free ("x", T) $ Free ("y", T')))))))
+ (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
+ ||> Sign.parent_path;
+
+
+ (* prove characteristic equations for primrec combinators *)
+
+ val _ = Old_Datatype_Aux.message config
+ "Proving characteristic theorems for primrec combinators ...";
+
+ val rec_thms =
+ map (fn t =>
+ Goal.prove_sorry_global thy2 [] [] t
+ (fn {context = ctxt, ...} => EVERY
+ [rewrite_goals_tac ctxt reccomb_defs,
+ rtac @{thm 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)]))
+ (Old_Datatype_Prop.make_primrecs reccomb_names descr thy2);
+ in
+ thy2
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> Global_Theory.note_thmss ""
+ [((Binding.name "rec", [Named_Theorems.add @{named_theorems nitpick_simp}]),
+ [(rec_thms, [])])]
+ ||> Sign.parent_path
+ |-> (fn thms => pair (reccomb_names, maps #2 thms))
+ end;
+
+
+(* case combinators *)
+
+fun prove_case_thms (config : Old_Datatype_Aux.config)
+ new_type_names descr reccomb_names primrec_thms thy =
+ let
+ val _ = Old_Datatype_Aux.message config
+ "Proving characteristic theorems for case combinators ...";
+
+ val ctxt = Proof_Context.init_global thy;
+ val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
+
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val used = fold Term.add_tfree_namesT recTs [];
+ val newTs = take (length (hd descr)) recTs;
+ val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
+
+ fun mk_dummyT dt = binder_types (Old_Datatype_Aux.typ_of_dtyp descr' dt) ---> T';
+
+ val case_dummy_fns =
+ map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val Ts' = map mk_dummyT (filter Old_Datatype_Aux.is_rec_type cargs)
+ in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
+
+ val case_names0 = map (fn s => Sign.full_bname thy1 ("case_" ^ s)) new_type_names;
+
+ (* define case combinators via primrec combinators *)
+
+ fun def_case ((((i, (_, _, constrs)), T as Type (Tcon, _)), name), recname) (defs, thy) =
+ if is_some (Ctr_Sugar.ctr_sugar_of ctxt Tcon) then
+ (defs, thy)
+ else
+ let
+ val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
+ let
+ val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
+ val Ts' = Ts @ map mk_dummyT (filter Old_Datatype_Aux.is_rec_type cargs);
+ val frees' = map2 (Old_Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
+ val frees = take (length cargs) frees';
+ val free = Old_Datatype_Aux.mk_Free "f" (Ts ---> T') j;
+ in
+ (free, fold_rev (absfree o 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 = flat (take i case_dummy_fns) @ fns2 @ flat (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 (Thm.def_name (Long_Name.base_name name)),
+ Logic.mk_equals (Const (name, caseT),
+ fold_rev lambda fns1
+ (list_comb (reccomb,
+ flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.declare_const_global decl |> snd
+ |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
+ in (defs @ [def_thm], thy') end;
+
+ val (case_defs, thy2) =
+ fold def_case (hd descr ~~ newTs ~~ case_names0 ~~ take (length newTs) reccomb_names)
+ ([], thy1);
+
+ fun prove_case t =
+ Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
+ EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]);
+
+ fun prove_cases (Type (Tcon, _)) ts =
+ (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
+ SOME {case_thms, ...} => case_thms
+ | NONE => map prove_case ts);
+
+ val case_thms =
+ map2 prove_cases newTs (Old_Datatype_Prop.make_cases case_names0 descr thy2);
+
+ fun case_name_of (th :: _) =
+ fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
+
+ val case_names = map case_name_of case_thms;
+ in
+ thy2
+ |> Context.theory_map
+ ((fold o fold) (Named_Theorems.add_thm @{named_theorems nitpick_simp}) case_thms)
+ |> Sign.parent_path
+ |> Old_Datatype_Aux.store_thmss "case" new_type_names case_thms
+ |-> (fn thmss => pair (thmss, case_names))
+ end;
+
+
+(* case splitting *)
+
+fun prove_split_thms (config : Old_Datatype_Aux.config)
+ new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
+ let
+ val _ = Old_Datatype_Aux.message config "Proving equations for case splitting ...";
+
+ val descr' = flat descr;
+ val recTs = Old_Datatype_Aux.get_rec_types descr';
+ val newTs = 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;
+ fun tac ctxt =
+ EVERY [rtac exhaustion' 1,
+ ALLGOALS (asm_simp_tac
+ (put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))];
+ in
+ (Goal.prove_sorry_global thy [] [] t1 (tac o #context),
+ Goal.prove_sorry_global thy [] [] t2 (tac o #context))
+ end;
+
+ val split_thm_pairs =
+ map prove_split_thms
+ (Old_Datatype_Prop.make_splits case_names descr thy ~~ constr_inject ~~
+ dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
+
+ val (split_thms, split_asm_thms) = split_list split_thm_pairs
+
+ in
+ thy
+ |> Old_Datatype_Aux.store_thms "split" new_type_names split_thms
+ ||>> Old_Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
+ |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
+ end;
+
+fun prove_case_cong_weaks new_type_names case_names descr thy =
+ let
+ fun prove_case_cong_weak t =
+ Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
+
+ val case_cong_weaks =
+ map prove_case_cong_weak (Old_Datatype_Prop.make_case_cong_weaks case_names descr thy);
+
+ in thy |> Old_Datatype_Aux.store_thms "case_cong_weak" new_type_names case_cong_weaks end;
+
+
+(* additional theorems for TFL *)
+
+fun prove_nchotomys (config : Old_Datatype_Aux.config) new_type_names descr casedist_thms thy =
+ let
+ val _ = Old_Datatype_Aux.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 ctxt i 0 =
+ EVERY [TRY (rtac disjI1 i), hyp_subst_tac ctxt i, REPEAT (rtac exI i), rtac refl i]
+ | tac ctxt i n = rtac disjI2 i THEN tac ctxt i (n - 1);
+ in
+ Goal.prove_sorry_global thy [] [] t
+ (fn {context = ctxt, ...} =>
+ EVERY [rtac allI 1,
+ Old_Datatype_Aux.exh_tac (K exhaustion) 1,
+ ALLGOALS (fn i => tac ctxt i (i - 1))])
+ end;
+
+ val nchotomys =
+ map prove_nchotomy (Old_Datatype_Prop.make_nchotomys descr ~~ casedist_thms);
+
+ in thy |> Old_Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
+
+fun prove_case_congs new_type_names case_names descr nchotomys case_thms thy =
+ let
+ fun prove_case_cong ((t, nchotomy), case_rewrites) =
+ let
+ val Const (@{const_name Pure.imp}, _) $ tm $ _ = t;
+ val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ 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
+ Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {context = ctxt, prems, ...} =>
+ let
+ val simplify = asm_simp_tac (put_simpset HOL_ss ctxt addsimps (prems @ case_rewrites))
+ in
+ EVERY [
+ simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1,
+ cut_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
+ (Old_Datatype_Prop.make_case_congs case_names descr thy ~~ nchotomys ~~ case_thms);
+
+ in thy |> Old_Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
+
+
+
+(** derive datatype props **)
+
+local
+
+fun make_dt_info descr induct inducts rec_names rec_rewrites
+ (index, (((((((((((_, (tname, _, _))), inject), distinct),
+ exhaust), nchotomy), case_name), case_rewrites), case_cong), case_cong_weak),
+ (split, split_asm))) =
+ (tname,
+ {index = index,
+ descr = descr,
+ inject = inject,
+ distinct = distinct,
+ induct = induct,
+ inducts = inducts,
+ exhaust = exhaust,
+ nchotomy = nchotomy,
+ rec_names = rec_names,
+ rec_rewrites = rec_rewrites,
+ case_name = case_name,
+ case_rewrites = case_rewrites,
+ case_cong = case_cong,
+ case_cong_weak = case_cong_weak,
+ split = split,
+ split_asm = split_asm});
+
+in
+
+fun derive_datatype_props config dt_names descr induct inject distinct thy2 =
+ let
+ val flat_descr = flat descr;
+ val new_type_names = map Long_Name.base_name dt_names;
+ val _ =
+ Old_Datatype_Aux.message config
+ ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
+
+ val (exhaust, thy3) = thy2
+ |> prove_casedist_thms config new_type_names descr induct
+ (Old_Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
+ val (nchotomys, thy4) = thy3
+ |> prove_nchotomys config new_type_names descr exhaust;
+ val ((rec_names, rec_rewrites), thy5) = thy4
+ |> prove_primrec_thms config new_type_names descr
+ (#inject o the o Symtab.lookup (Old_Datatype_Data.get_all thy4)) inject
+ (distinct,
+ Old_Datatype_Data.all_distincts thy2 (Old_Datatype_Aux.get_rec_types flat_descr)) induct;
+ val ((case_rewrites, case_names), thy6) = thy5
+ |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
+ val (case_congs, thy7) = thy6
+ |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
+ val (case_cong_weaks, thy8) = thy7
+ |> prove_case_cong_weaks new_type_names case_names descr;
+ val (splits, thy9) = thy8
+ |> prove_split_thms config new_type_names case_names descr
+ inject distinct exhaust case_rewrites;
+
+ val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
+ val dt_infos =
+ map_index
+ (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
+ (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
+ case_names ~~ case_rewrites ~~ case_congs ~~ case_cong_weaks ~~ splits);
+ val dt_names = map fst dt_infos;
+ val prfx = Binding.qualify true (space_implode "_" new_type_names);
+ val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
+ val named_rules = flat (map_index (fn (i, tname) =>
+ [((Binding.empty, [Induct.induct_type tname]), [([nth inducts i], [])]),
+ ((Binding.empty, [Induct.cases_type tname]), [([nth exhaust i], [])])]) dt_names);
+ val unnamed_rules = map (fn induct =>
+ ((Binding.empty, [Rule_Cases.inner_rule, Induct.induct_type ""]), [([induct], [])]))
+ (drop (length dt_names) inducts);
+
+ val ctxt = Proof_Context.init_global thy9;
+ val case_combs =
+ map (Proof_Context.read_const {proper = true, strict = true} ctxt) case_names;
+ val constrss = map (fn (dtname, {descr, index, ...}) =>
+ map (Proof_Context.read_const {proper = true, strict = true} ctxt o fst)
+ (#3 (the (AList.lookup op = descr index)))) dt_infos;
+ in
+ thy9
+ |> Global_Theory.note_thmss ""
+ ([((prfx (Binding.name "simps"), []), [(simps, [])]),
+ ((prfx (Binding.name "inducts"), []), [(inducts, [])]),
+ ((prfx (Binding.name "splits"), []), [(maps (fn (x, y) => [x, y]) splits, [])]),
+ ((Binding.empty, [Simplifier.simp_add]),
+ [(flat case_rewrites @ flat distinct @ rec_rewrites, [])]),
+ ((Binding.empty, [Code.add_default_eqn_attribute]), [(rec_rewrites, [])]),
+ ((Binding.empty, [iff_add]), [(flat inject, [])]),
+ ((Binding.empty, [Classical.safe_elim NONE]),
+ [(map (fn th => th RS notE) (flat distinct), [])]),
+ ((Binding.empty, [Simplifier.cong_add]), [(case_cong_weaks, [])]),
+ ((Binding.empty, [Induct.induct_simp_add]), [(flat (distinct @ inject), [])])] @
+ named_rules @ unnamed_rules)
+ |> snd
+ |> Old_Datatype_Data.register dt_infos
+ |> Context.theory_map (fold2 Case_Translation.register case_combs constrss)
+ |> Old_Datatype_Data.interpretation_data (config, dt_names)
+ |> pair dt_names
+ end;
+
+end;
+
+
+
+(** declare existing type as datatype **)
+
+local
+
+fun prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct thy1 =
+ let
+ val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
+ val new_type_names = map Long_Name.base_name dt_names;
+ val prfx = Binding.qualify true (space_implode "_" new_type_names);
+ val (((inject, distinct), [(_, [induct])]), thy2) =
+ thy1
+ |> Old_Datatype_Aux.store_thmss "inject" new_type_names raw_inject
+ ||>> Old_Datatype_Aux.store_thmss "distinct" new_type_names raw_distinct
+ ||>> Global_Theory.note_thmss ""
+ [((prfx (Binding.name "induct"), [Old_Datatype_Data.mk_case_names_induct descr]),
+ [([raw_induct], [])])];
+ in
+ thy2
+ |> derive_datatype_props config dt_names [descr] induct inject distinct
+ end;
+
+fun gen_rep_datatype prep_term config after_qed raw_ts thy =
+ let
+ val ctxt = Proof_Context.init_global thy;
+
+ fun constr_of_term (Const (c, T)) = (c, T)
+ | constr_of_term t = error ("Not a constant: " ^ Syntax.string_of_term ctxt t);
+ fun no_constr (c, T) =
+ error ("Bad constructor: " ^ Proof_Context.markup_const ctxt c ^ "::" ^
+ Syntax.string_of_typ ctxt T);
+ fun type_of_constr (cT as (_, T)) =
+ let
+ val frees = Term.add_tfreesT T [];
+ val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_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 (Old_Datatype_Data.get_all thy) tyco then SOME tyco else NONE) raw_cs of
+ [] => ()
+ | tycos => error ("Type(s) " ^ commas_quote tycos ^ " already represented inductively"));
+ 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
+ |> foldr1 (Sorts.inter_sort (Sign.classes_of thy));
+ val sorts = map inter_sort ms;
+ val vs = Name.invent_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 (Old_Datatype_Aux.dtyp_of_typ (map (rpair vs o fst) cs)) o binder_types;
+ val dt_names = map fst cs;
+
+ fun mk_spec (i, (tyco, constr)) =
+ (i, (tyco, map Old_Datatype_Aux.DtTFree vs, (map o apsnd) dtyps_of_typ constr));
+ val descr = map_index mk_spec cs;
+ val injs = Old_Datatype_Prop.make_injs [descr];
+ val half_distincts = Old_Datatype_Prop.make_distincts [descr];
+ val ind = Old_Datatype_Prop.make_ind [descr];
+ val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
+
+ fun after_qed' raw_thms =
+ let
+ val [[[raw_induct]], raw_inject, half_distinct] =
+ unflat rules (map Drule.zero_var_indexes_list raw_thms);
+ (*FIXME somehow dubious*)
+ in
+ Proof_Context.background_theory_result (* FIXME !? *)
+ (prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct)
+ #-> after_qed
+ end;
+ in
+ ctxt
+ |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
+ end;
+
+in
+
+val rep_datatype = gen_rep_datatype Sign.cert_term;
+val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global;
+
+end;
+
+
+(* outer syntax *)
+
+val _ =
+ Outer_Syntax.command @{command_spec "rep_datatype"} "represent existing types inductively"
+ (Scan.repeat1 Parse.term >> (fn ts =>
+ Toplevel.theory_to_proof (rep_datatype_cmd Old_Datatype_Aux.default_config (K I) ts)));
+
+end;
--- a/src/HOL/Tools/Quickcheck/abstract_generators.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/abstract_generators.ML Mon Sep 01 16:17:46 2014 +0200
@@ -38,19 +38,19 @@
val vs = map dest_TFree (snd (dest_Type (body_type (fastype_of (hd constrs)))))
val name = Long_Name.base_name tyco
fun mk_dconstrs (Const (s, T)) =
- (s, map (Datatype_Aux.dtyp_of_typ [(tyco, vs)]) (binder_types T))
+ (s, map (Old_Datatype_Aux.dtyp_of_typ [(tyco, vs)]) (binder_types T))
| mk_dconstrs t = error (Syntax.string_of_term ctxt t ^
" is not a valid constructor for quickcheck_generator, which expects a constant.")
fun the_descr _ _ =
let
- val descr = [(0, (tyco, map Datatype_Aux.DtTFree vs, map mk_dconstrs constrs))]
+ val descr = [(0, (tyco, map Old_Datatype_Aux.DtTFree vs, map mk_dconstrs constrs))]
in
(descr, vs, [tyco], name, ([name], []), ([Type (tyco, map TFree vs)], []))
end
fun ensure_sort (sort, instantiate) =
Quickcheck_Common.ensure_sort (((@{sort typerep}, @{sort term_of}), sort),
(the_descr, instantiate))
- Datatype_Aux.default_config [tyco]
+ Old_Datatype_Aux.default_config [tyco]
in
thy
|> ensure_sort (@{sort full_exhaustive}, Exhaustive_Generators.instantiate_full_exhaustive_datatype)
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Mon Sep 01 16:17:46 2014 +0200
@@ -23,9 +23,9 @@
val setup_bounded_forall_datatype_interpretation : theory -> theory
val setup: theory -> theory
- val instantiate_full_exhaustive_datatype : Datatype_Aux.config -> Datatype_Aux.descr ->
+ val instantiate_full_exhaustive_datatype : Old_Datatype_Aux.config -> Old_Datatype_Aux.descr ->
(string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
- val instantiate_exhaustive_datatype : Datatype_Aux.config -> Datatype_Aux.descr ->
+ val instantiate_exhaustive_datatype : Old_Datatype_Aux.config -> Old_Datatype_Aux.descr ->
(string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
end;
@@ -95,7 +95,7 @@
let
val (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, exhaustives) = generics
val rhss =
- Datatype_Aux.interpret_construction descr vs
+ Old_Datatype_Aux.interpret_construction descr vs
{ atyp = mk_call, dtyp = mk_aux_call }
|> (map o apfst) Type
|> map (fn (T, cs) => map (mk_consexpr T) cs)
@@ -198,13 +198,13 @@
fun contains_recursive_type_under_function_types xs =
exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
- (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype_Aux.DtRec _) => true | _ => false)))) xs
+ (case Old_Datatype_Aux.strip_dtyp dT of (_ :: _, Old_Datatype_Aux.DtRec _) => true | _ => false)))) xs
fun instantiate_datatype (name, constprfx, sort, mk_equations, mk_T, argnames)
config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
if not (contains_recursive_type_under_function_types descr) then
let
- val _ = Datatype_Aux.message config ("Creating " ^ name ^ "...")
+ val _ = Old_Datatype_Aux.message config ("Creating " ^ name ^ "...")
val fullnames = map (prefix (constprfx ^ "_")) (names @ auxnames)
in
thy
@@ -215,7 +215,7 @@
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
end
else
- (Datatype_Aux.message config
+ (Old_Datatype_Aux.message config
("Creation of " ^ name ^ " failed because the datatype is recursive under a function type");
thy)
@@ -254,7 +254,7 @@
let
val cnstrs = flat (maps
(map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
- (Symtab.dest (Datatype_Data.get_all (Proof_Context.theory_of ctxt))))
+ (Symtab.dest (Old_Datatype_Data.get_all (Proof_Context.theory_of ctxt))))
fun is_constrt (Const (s, T), ts) = (case (AList.lookup (op =) cnstrs s, body_type T) of
(SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname'
| _ => false)
@@ -537,19 +537,19 @@
Quickcheck_Common.datatype_interpretation (@{sort exhaustive}, instantiate_exhaustive_datatype)
val setup_bounded_forall_datatype_interpretation =
- Datatype_Data.interpretation (Quickcheck_Common.ensure_sort
+ Old_Datatype_Data.interpretation (Quickcheck_Common.ensure_sort
(((@{sort type}, @{sort type}), @{sort bounded_forall}),
- (Datatype_Data.the_descr, instantiate_bounded_forall_datatype)))
+ (Old_Datatype_Data.the_descr, instantiate_bounded_forall_datatype)))
val active = Attrib.setup_config_bool @{binding quickcheck_exhaustive_active} (K true);
val setup =
Quickcheck_Common.datatype_interpretation (@{sort full_exhaustive}, instantiate_full_exhaustive_datatype)
-(* #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
+(* #> Old_Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
(((@{sort typerep}, @{sort term_of}), @{sort exhaustive}), instantiate_exhaustive_datatype))
- #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
+ #> Old_Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
(((@{sort typerep}, @{sort term_of}), @{sort fast_exhaustive}), instantiate_fast_exhaustive_datatype))
- #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
+ #> Old_Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
(((@{sort type}, @{sort type}), @{sort bounded_forall}), instantiate_bounded_forall_datatype))*)
#> Context.theory_map (Quickcheck.add_tester ("exhaustive", (active, test_goals)))
#> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs))
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Mon Sep 01 16:17:46 2014 +0200
@@ -162,7 +162,7 @@
in snd (fold mk_apply xs (Ts ---> simpleT, mk_cons c (Ts ---> simpleT))) end
fun mk_rhs exprs = foldr1 mk_sum exprs
val rhss =
- Datatype_Aux.interpret_construction descr vs
+ Old_Datatype_Aux.interpret_construction descr vs
{ atyp = mk_call, dtyp = mk_aux_call }
|> (map o apfst) Type
|> map (fn (T, cs) => map (mk_consexpr T) cs)
@@ -175,11 +175,11 @@
fun contains_recursive_type_under_function_types xs =
exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
- (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype_Aux.DtRec _) => true | _ => false)))) xs
+ (case Old_Datatype_Aux.strip_dtyp dT of (_ :: _, Old_Datatype_Aux.DtRec _) => true | _ => false)))) xs
fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
let
- val _ = Datatype_Aux.message config "Creating narrowing generators ...";
+ val _ = Old_Datatype_Aux.message config "Creating narrowing generators ...";
val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames);
in
if not (contains_recursive_type_under_function_types descr) then
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML Mon Sep 01 16:17:46 2014 +0200
@@ -23,15 +23,15 @@
val generator_test_goal_terms :
generator -> Proof.context -> bool -> (string * typ) list
-> (term * term list) list -> Quickcheck.result list
- type instantiation = Datatype_Aux.config -> Datatype_Aux.descr -> (string * sort) list
+ type instantiation = Old_Datatype_Aux.config -> Old_Datatype_Aux.descr -> (string * sort) list
-> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
val ensure_sort :
(((sort * sort) * sort) *
- ((theory -> string list -> Datatype_Aux.descr * (string * sort) list * string list
+ ((theory -> string list -> Old_Datatype_Aux.descr * (string * sort) list * string list
* string * (string list * string list) * (typ list * typ list)) * instantiation))
- -> Datatype_Aux.config -> string list -> theory -> theory
+ -> Old_Datatype_Aux.config -> string list -> theory -> theory
val ensure_common_sort_datatype :
- (sort * instantiation) -> Datatype_Aux.config -> string list -> theory -> theory
+ (sort * instantiation) -> Old_Datatype_Aux.config -> string list -> theory -> theory
val datatype_interpretation : (sort * instantiation) -> theory -> theory
val gen_mk_parametric_generator_expr :
(((Proof.context -> term * term list -> term) * term) * typ)
@@ -387,7 +387,7 @@
(** ensuring sort constraints **)
-type instantiation = Datatype_Aux.config -> Datatype_Aux.descr -> (string * sort) list
+type instantiation = Old_Datatype_Aux.config -> Old_Datatype_Aux.descr -> (string * sort) list
-> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
fun perhaps_constrain thy insts raw_vs =
@@ -406,7 +406,7 @@
val (descr, raw_vs, tycos, prfx, (names, auxnames), raw_TUs) = the_descr thy raw_tycos
val vs = (map o apsnd) (curry (Sorts.inter_sort algebra) sort_vs) raw_vs
fun insts_of sort constr = (map (rpair sort) o flat o maps snd o maps snd)
- (Datatype_Aux.interpret_construction descr vs constr)
+ (Old_Datatype_Aux.interpret_construction descr vs constr)
val insts = insts_of sort { atyp = single, dtyp = (K o K o K) [] }
@ insts_of aux_sort { atyp = K [], dtyp = K o K }
val has_inst = exists (fn tyco => Sorts.has_instance algebra tyco sort) tycos;
@@ -419,9 +419,9 @@
end;
fun ensure_common_sort_datatype (sort, instantiate) =
- ensure_sort (((@{sort typerep}, @{sort term_of}), sort), (Datatype_Data.the_descr, instantiate))
+ ensure_sort (((@{sort typerep}, @{sort term_of}), sort), (Old_Datatype_Data.the_descr, instantiate))
-val datatype_interpretation = Datatype_Data.interpretation o ensure_common_sort_datatype
+val datatype_interpretation = Old_Datatype_Data.interpretation o ensure_common_sort_datatype
(** generic parametric compilation **)
--- a/src/HOL/Tools/Quickcheck/random_generators.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML Mon Sep 01 16:17:46 2014 +0200
@@ -16,7 +16,7 @@
-> Proof.context -> Proof.context
val put_counterexample_report: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed)
-> Proof.context -> Proof.context
- val instantiate_random_datatype : Datatype_Aux.config -> Datatype_Aux.descr ->
+ val instantiate_random_datatype : Old_Datatype_Aux.config -> Old_Datatype_Aux.descr ->
(string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
val setup: theory -> theory
end;
@@ -97,7 +97,7 @@
val eqs0 = [subst_v @{term "0::natural"} eq,
subst_v (@{const Code_Numeral.Suc} $ t_k) eq];
val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0;
- val ((_, (_, eqs2)), lthy') = Primrec.add_primrec_simple
+ val ((_, (_, eqs2)), lthy') = Old_Primrec.add_primrec_simple
[((Binding.conceal (Binding.name random_aux), T), NoSyn)] eqs1 lthy;
val cT_random_aux = inst pt_random_aux;
val cT_rhs = inst pt_rhs;
@@ -204,10 +204,10 @@
mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $
mk_random_fun_lift fTs t;
val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size');
- val size = Option.map snd (Datatype_Aux.find_shortest_path descr k)
+ val size = Option.map snd (Old_Datatype_Aux.find_shortest_path descr k)
|> the_default 0;
in (SOME size, (t, fTs ---> T)) end;
- val tss = Datatype_Aux.interpret_construction descr vs
+ val tss = Old_Datatype_Aux.interpret_construction descr vs
{ atyp = mk_random_call, dtyp = mk_random_aux_call };
fun mk_consexpr simpleT (c, xs) =
let
@@ -245,9 +245,9 @@
fun instantiate_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
let
- val _ = Datatype_Aux.message config "Creating quickcheck generators ...";
+ val _ = Old_Datatype_Aux.message config "Creating quickcheck generators ...";
val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
- fun mk_size_arg k = case Datatype_Aux.find_shortest_path descr k
+ fun mk_size_arg k = case Old_Datatype_Aux.find_shortest_path descr k
of SOME (_, l) => if l = 0 then size
else @{term "max :: natural \<Rightarrow> natural \<Rightarrow> natural"}
$ HOLogic.mk_number @{typ natural} l $ size
--- a/src/HOL/Tools/SMT/smt_normalize.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Mon Sep 01 16:17:46 2014 +0200
@@ -144,7 +144,7 @@
| @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
| _ => raise CTERM ("no equation", [ct]))
- fun get_constrs thy (Type (n, _)) = these (Datatype_Data.get_constrs thy n)
+ fun get_constrs thy (Type (n, _)) = these (Old_Datatype_Data.get_constrs thy n)
| get_constrs _ _ = []
fun is_constr thy (n, T) =
--- a/src/HOL/Tools/TFL/casesplit.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/TFL/casesplit.ML Mon Sep 01 16:17:46 2014 +0200
@@ -24,7 +24,7 @@
Type(ty_str, _) => ty_str
| TFree(s,_) => error ("Free type: " ^ s)
| TVar((s,i),_) => error ("Free variable: " ^ s)
- val {induct, ...} = Datatype_Data.the_info thy ty_str
+ val {induct, ...} = Old_Datatype_Data.the_info thy ty_str
in
cases_thm_of_induct_thm induct
end;
--- a/src/HOL/Tools/TFL/tfl.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/TFL/tfl.ML Mon Sep 01 16:17:46 2014 +0200
@@ -435,7 +435,7 @@
put_simpset HOL_basic_ss ctxt
addsimps case_rewrites
|> fold (Simplifier.add_cong o #case_cong_weak o snd)
- (Symtab.dest (Datatype_Data.get_all theory))
+ (Symtab.dest (Old_Datatype_Data.get_all theory))
val corollaries' = map (Simplifier.simplify case_simpset) corollaries
val extract = Rules.CONTEXT_REWRITE_RULE
(f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs)
--- a/src/HOL/Tools/TFL/thry.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/TFL/thry.ML Mon Sep 01 16:17:46 2014 +0200
@@ -58,20 +58,20 @@
*---------------------------------------------------------------------------*)
fun match_info thy dtco =
- case (Datatype_Data.get_info thy dtco,
- Datatype_Data.get_constrs thy dtco) of
+ case (Old_Datatype_Data.get_info thy dtco,
+ Old_Datatype_Data.get_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 Datatype_Data.get_info thy dtco of
+fun induct_info thy dtco = case Old_Datatype_Data.get_info thy dtco of
NONE => NONE
| SOME {nchotomy, ...} =>
SOME {nchotomy = nchotomy,
- constructors = (map Const o the o Datatype_Data.get_constrs thy) dtco};
+ constructors = (map Const o the o Old_Datatype_Data.get_constrs thy) dtco};
fun extract_info thy =
- let val infos = map snd (Symtab.dest (Datatype_Data.get_all thy))
+ let val infos = map snd (Symtab.dest (Old_Datatype_Data.get_all thy))
in {case_congs = map (mk_meta_eq o #case_cong) infos,
case_rewrites = maps (map mk_meta_eq o #case_rewrites) infos}
end;
--- a/src/HOL/Tools/inductive_realizer.ML Mon Sep 01 16:17:46 2014 +0200
+++ b/src/HOL/Tools/inductive_realizer.ML Mon Sep 01 16:17:46 2014 +0200
@@ -247,7 +247,7 @@
thy
|> f (map snd dts)
|-> (fn dtinfo => pair (map fst dts, SOME dtinfo))
- handle Datatype_Aux.Datatype_Empty name' =>
+ handle Old_Datatype_Aux.Datatype_Empty name' =>
let
val name = Long_Name.base_name name';
val dname = singleton (Name.variant_list used) "Dummy";
@@ -308,15 +308,15 @@
val ((dummies, some_dt_names), thy2) =
thy1
- |> add_dummies (Datatype.add_datatype {strict = false, quiet = false})
+ |> add_dummies (Old_Datatype.add_datatype {strict = false, quiet = false})
(map (pair false) dts) []
||> Extraction.add_typeof_eqns_i ty_eqs
||> Extraction.add_realizes_eqns_i rlz_eqs;
val dt_names = these some_dt_names;
- val case_thms = map (#case_rewrites o Datatype_Data.the_info thy2) dt_names;
+ val case_thms = map (#case_rewrites o Old_Datatype_Data.the_info thy2) dt_names;
val rec_thms =
if null dt_names then []
- else #rec_rewrites (Datatype_Data.the_info thy2 (hd dt_names));
+ else #rec_rewrites (Old_Datatype_Data.the_info thy2 (hd dt_names));
val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o
HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) rec_thms);
val (constrss, _) = fold_map (fn (s, rs) => fn (recs, dummies) =>
@@ -402,7 +402,7 @@
val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
(Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy;
val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp)))
- (Datatype_Aux.split_conj_thm thm');
+ (Old_Datatype_Aux.split_conj_thm thm');
val ([thms'], thy'') = Global_Theory.add_thmss
[((Binding.qualified_name (space_implode "_"
(Long_Name.qualify qualifier "inducts" :: vs' @ Ps @