diff --git a/src/HOL/Decision_Procs/Commutative_Ring.thy b/src/HOL/Decision_Procs/Commutative_Ring.thy --- a/src/HOL/Decision_Procs/Commutative_Ring.thy +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy @@ -1,963 +1,963 @@ (* Title: HOL/Decision_Procs/Commutative_Ring.thy Author: Bernhard Haeupler, Stefan Berghofer, and Amine Chaieb Proving equalities in commutative rings done "right" in Isabelle/HOL. *) section \Proving equalities in commutative rings\ theory Commutative_Ring imports Conversions Algebra_Aux "HOL-Library.Code_Target_Numeral" begin text \Syntax of multivariate polynomials (pol) and polynomial expressions.\ datatype pol = Pc int | Pinj nat pol | PX pol nat pol datatype polex = Var nat | Const int | Add polex polex | Sub polex polex | Mul polex polex | Pow polex nat | Neg polex text \Interpretation functions for the shadow syntax.\ context cring begin definition in_carrier :: "'a list \ bool" where "in_carrier xs \ (\x\set xs. x \ carrier R)" lemma in_carrier_Nil: "in_carrier []" by (simp add: in_carrier_def) lemma in_carrier_Cons: "x \ carrier R \ in_carrier xs \ in_carrier (x # xs)" by (simp add: in_carrier_def) lemma drop_in_carrier [simp]: "in_carrier xs \ in_carrier (drop n xs)" using set_drop_subset [of n xs] by (auto simp add: in_carrier_def) primrec head :: "'a list \ 'a" where "head [] = \" | "head (x # xs) = x" lemma head_closed [simp]: "in_carrier xs \ head xs \ carrier R" by (cases xs) (simp_all add: in_carrier_def) primrec Ipol :: "'a list \ pol \ 'a" where "Ipol l (Pc c) = \c\" | "Ipol l (Pinj i P) = Ipol (drop i l) P" | "Ipol l (PX P x Q) = Ipol l P \ head l [^] x \ Ipol (drop 1 l) Q" lemma Ipol_Pc: "Ipol l (Pc 0) = \" "Ipol l (Pc 1) = \" "Ipol l (Pc (numeral n)) = \numeral n\" "Ipol l (Pc (- numeral n)) = \ \numeral n\" by simp_all lemma Ipol_closed [simp]: "in_carrier l \ Ipol l p \ carrier R" by (induct p arbitrary: l) simp_all primrec Ipolex :: "'a list \ polex \ 'a" where "Ipolex l (Var n) = head (drop n l)" | "Ipolex l (Const i) = \i\" | "Ipolex l (Add P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Sub P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Mul P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Pow p n) = Ipolex l p [^] n" | "Ipolex l (Neg P) = \ Ipolex l P" lemma Ipolex_Const: "Ipolex l (Const 0) = \" "Ipolex l (Const 1) = \" "Ipolex l (Const (numeral n)) = \numeral n\" by simp_all end text \Create polynomial normalized polynomials given normalized inputs.\ definition mkPinj :: "nat \ pol \ pol" where "mkPinj x P = (case P of Pc c \ Pc c | Pinj y P \ Pinj (x + y) P | PX p1 y p2 \ Pinj x P)" definition mkPX :: "pol \ nat \ pol \ pol" where "mkPX P i Q = (case P of Pc c \ if c = 0 then mkPinj 1 Q else PX P i Q | Pinj j R \ PX P i Q | PX P2 i2 Q2 \ if Q2 = Pc 0 then PX P2 (i + i2) Q else PX P i Q)" text \Defining the basic ring operations on normalized polynomials\ function add :: "pol \ pol \ pol" (infixl "\+\" 65) where "Pc a \+\ Pc b = Pc (a + b)" | "Pc c \+\ Pinj i P = Pinj i (P \+\ Pc c)" | "Pinj i P \+\ Pc c = Pinj i (P \+\ Pc c)" | "Pc c \+\ PX P i Q = PX P i (Q \+\ Pc c)" | "PX P i Q \+\ Pc c = PX P i (Q \+\ Pc c)" | "Pinj x P \+\ Pinj y Q = (if x = y then mkPinj x (P \+\ Q) else (if x > y then mkPinj y (Pinj (x - y) P \+\ Q) else mkPinj x (Pinj (y - x) Q \+\ P)))" | "Pinj x P \+\ PX Q y R = (if x = 0 then P \+\ PX Q y R else (if x = 1 then PX Q y (R \+\ P) else PX Q y (R \+\ Pinj (x - 1) P)))" | "PX P x R \+\ Pinj y Q = (if y = 0 then PX P x R \+\ Q else (if y = 1 then PX P x (R \+\ Q) else PX P x (R \+\ Pinj (y - 1) Q)))" | "PX P1 x P2 \+\ PX Q1 y Q2 = (if x = y then mkPX (P1 \+\ Q1) x (P2 \+\ Q2) else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \+\ Q1) y (P2 \+\ Q2) else mkPX (PX Q1 (y - x) (Pc 0) \+\ P1) x (P2 \+\ Q2)))" by pat_completeness auto termination by (relation "measure (\(x, y). size x + size y)") auto function mul :: "pol \ pol \ pol" (infixl "\*\" 70) where "Pc a \*\ Pc b = Pc (a * b)" | "Pc c \*\ Pinj i P = (if c = 0 then Pc 0 else mkPinj i (P \*\ Pc c))" | "Pinj i P \*\ Pc c = (if c = 0 then Pc 0 else mkPinj i (P \*\ Pc c))" | "Pc c \*\ PX P i Q = (if c = 0 then Pc 0 else mkPX (P \*\ Pc c) i (Q \*\ Pc c))" | "PX P i Q \*\ Pc c = (if c = 0 then Pc 0 else mkPX (P \*\ Pc c) i (Q \*\ Pc c))" | "Pinj x P \*\ Pinj y Q = (if x = y then mkPinj x (P \*\ Q) else (if x > y then mkPinj y (Pinj (x - y) P \*\ Q) else mkPinj x (Pinj (y - x) Q \*\ P)))" | "Pinj x P \*\ PX Q y R = (if x = 0 then P \*\ PX Q y R else (if x = 1 then mkPX (Pinj x P \*\ Q) y (R \*\ P) else mkPX (Pinj x P \*\ Q) y (R \*\ Pinj (x - 1) P)))" | "PX P x R \*\ Pinj y Q = (if y = 0 then PX P x R \*\ Q else (if y = 1 then mkPX (Pinj y Q \*\ P) x (R \*\ Q) else mkPX (Pinj y Q \*\ P) x (R \*\ Pinj (y - 1) Q)))" | "PX P1 x P2 \*\ PX Q1 y Q2 = mkPX (P1 \*\ Q1) (x + y) (P2 \*\ Q2) \+\ (mkPX (P1 \*\ mkPinj 1 Q2) x (Pc 0) \+\ (mkPX (Q1 \*\ mkPinj 1 P2) y (Pc 0)))" by pat_completeness auto termination by (relation "measure (\(x, y). size x + size y)") (auto simp add: mkPinj_def split: pol.split) text \Negation\ primrec neg :: "pol \ pol" where "neg (Pc c) = Pc (- c)" | "neg (Pinj i P) = Pinj i (neg P)" | "neg (PX P x Q) = PX (neg P) x (neg Q)" text \Subtraction\ definition sub :: "pol \ pol \ pol" (infixl "\-\" 65) where "sub P Q = P \+\ neg Q" text \Square for Fast Exponentiation\ primrec sqr :: "pol \ pol" where "sqr (Pc c) = Pc (c * c)" | "sqr (Pinj i P) = mkPinj i (sqr P)" | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \+\ mkPX (Pc 2 \*\ A \*\ mkPinj 1 B) x (Pc 0)" text \Fast Exponentiation\ fun pow :: "nat \ pol \ pol" where pow_if [simp del]: "pow n P = (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) else P \*\ pow (n div 2) (sqr P))" lemma pow_simps [simp]: "pow 0 P = Pc 1" "pow (2 * n) P = pow n (sqr P)" "pow (Suc (2 * n)) P = P \*\ pow n (sqr P)" by (simp_all add: pow_if) lemma even_pow: "even n \ pow n P = pow (n div 2) (sqr P)" by (erule evenE) simp lemma odd_pow: "odd n \ pow n P = P \*\ pow (n div 2) (sqr P)" by (erule oddE) simp text \Normalization of polynomial expressions\ primrec norm :: "polex \ pol" where "norm (Var n) = (if n = 0 then PX (Pc 1) 1 (Pc 0) else Pinj n (PX (Pc 1) 1 (Pc 0)))" | "norm (Const i) = Pc i" | "norm (Add P Q) = norm P \+\ norm Q" | "norm (Sub P Q) = norm P \-\ norm Q" | "norm (Mul P Q) = norm P \*\ norm Q" | "norm (Pow P n) = pow n (norm P)" | "norm (Neg P) = neg (norm P)" context cring begin text \mkPinj preserve semantics\ lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" by (induct B) (auto simp add: mkPinj_def algebra_simps) text \mkPX preserves semantics\ lemma mkPX_ci: "in_carrier l \ Ipol l (mkPX A b C) = Ipol l (PX A b C)" by (cases A) (auto simp add: mkPX_def mkPinj_ci nat_pow_mult [symmetric] m_ac) text \Correctness theorems for the implemented operations\ text \Negation\ lemma neg_ci: "in_carrier l \ Ipol l (neg P) = \ (Ipol l P)" by (induct P arbitrary: l) (auto simp add: minus_add l_minus) text \Addition\ lemma add_ci: "in_carrier l \ Ipol l (P \+\ Q) = Ipol l P \ Ipol l Q" proof (induct P Q arbitrary: l rule: add.induct) case (6 x P y Q) consider "x < y" | "x = y" | "x > y" by arith then show ?case proof cases case 1 with 6 show ?thesis by (simp add: mkPinj_ci a_ac) next case 2 with 6 show ?thesis by (simp add: mkPinj_ci) next case 3 with 6 show ?thesis by (simp add: mkPinj_ci) qed next case (7 x P Q y R) consider "x = 0" | "x = 1" | "x > 1" by arith then show ?case proof cases case 1 with 7 show ?thesis by simp next case 2 with 7 show ?thesis by (simp add: a_ac) next case 3 with 7 show ?thesis by (cases x) (simp_all add: a_ac) qed next case (8 P x R y Q) then show ?case by (simp add: a_ac) next case (9 P1 x P2 Q1 y Q2) consider "x = y" | d where "d + x = y" | d where "d + y = x" by atomize_elim arith then show ?case proof cases case 1 with 9 show ?thesis by (simp add: mkPX_ci r_distr a_ac m_ac) next case 2 with 9 show ?thesis by (auto simp add: mkPX_ci nat_pow_mult [symmetric] r_distr a_ac m_ac) next case 3 with 9 show ?thesis by (auto simp add: nat_pow_mult [symmetric] mkPX_ci r_distr a_ac m_ac) qed qed (auto simp add: a_ac m_ac) text \Multiplication\ lemma mul_ci: "in_carrier l \ Ipol l (P \*\ Q) = Ipol l P \ Ipol l Q" by (induct P Q arbitrary: l rule: mul.induct) (simp_all add: mkPX_ci mkPinj_ci a_ac m_ac r_distr add_ci nat_pow_mult [symmetric]) text \Subtraction\ lemma sub_ci: "in_carrier l \ Ipol l (P \-\ Q) = Ipol l P \ Ipol l Q" by (simp add: add_ci neg_ci sub_def minus_eq) text \Square\ lemma sqr_ci: "in_carrier ls \ Ipol ls (sqr P) = Ipol ls P \ Ipol ls P" by (induct P arbitrary: ls) (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci l_distr r_distr a_ac m_ac nat_pow_mult [symmetric] of_int_2) text \Power\ lemma pow_ci: "in_carrier ls \ Ipol ls (pow n P) = Ipol ls P [^] n" proof (induct n arbitrary: P rule: less_induct) case (less k) consider "k = 0" | "k > 0" by arith then show ?case proof cases case 1 then show ?thesis by simp next case 2 then have "k div 2 < k" by arith with less have *: "Ipol ls (pow (k div 2) (sqr P)) = Ipol ls (sqr P) [^] (k div 2)" by simp show ?thesis proof (cases "even k") case True with * \in_carrier ls\ show ?thesis by (simp add: even_pow sqr_ci nat_pow_distrib nat_pow_mult mult_2 [symmetric] even_two_times_div_two) next case False with * \in_carrier ls\ show ?thesis by (simp add: odd_pow mul_ci sqr_ci nat_pow_distrib nat_pow_mult mult_2 [symmetric] trans [OF nat_pow_Suc m_comm, symmetric]) qed qed qed text \Normalization preserves semantics\ lemma norm_ci: "in_carrier l \ Ipolex l Pe = Ipol l (norm Pe)" by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) text \Reflection lemma: Key to the (incomplete) decision procedure\ lemma norm_eq: assumes "in_carrier l" and "norm P1 = norm P2" shows "Ipolex l P1 = Ipolex l P2" proof - from assms have "Ipol l (norm P1) = Ipol l (norm P2)" by simp with assms show ?thesis by (simp only: norm_ci) qed end text \Monomials\ datatype mon = Mc int | Minj nat mon | MX nat mon primrec (in cring) Imon :: "'a list \ mon \ 'a" where "Imon l (Mc c) = \c\" | "Imon l (Minj i M) = Imon (drop i l) M" | "Imon l (MX x M) = Imon (drop 1 l) M \ head l [^] x" lemma (in cring) Imon_closed [simp]: "in_carrier l \ Imon l m \ carrier R" by (induct m arbitrary: l) simp_all definition mkMinj :: "nat \ mon \ mon" where "mkMinj i M = (case M of Mc c \ Mc c | Minj j M \ Minj (i + j) M | _ \ Minj i M)" definition Minj_pred :: "nat \ mon \ mon" where "Minj_pred i M = (if i = 1 then M else mkMinj (i - 1) M)" primrec mkMX :: "nat \ mon \ mon" where "mkMX i (Mc c) = MX i (Mc c)" | "mkMX i (Minj j M) = (if j = 0 then mkMX i M else MX i (Minj_pred j M))" | "mkMX i (MX j M) = MX (i + j) M" lemma (in cring) mkMinj_correct: "Imon l (mkMinj i M) = Imon l (Minj i M)" by (simp add: mkMinj_def add.commute split: mon.split) lemma (in cring) Minj_pred_correct: "0 < i \ Imon (drop 1 l) (Minj_pred i M) = Imon l (Minj i M)" by (simp add: Minj_pred_def mkMinj_correct) lemma (in cring) mkMX_correct: "in_carrier l \ Imon l (mkMX i M) = Imon l M \ head l [^] i" by (induct M) (simp_all add: Minj_pred_correct [simplified] nat_pow_mult [symmetric] m_ac split: mon.split) fun cfactor :: "pol \ int \ pol \ pol" where "cfactor (Pc c') c = (Pc (c' mod c), Pc (c' div c))" | "cfactor (Pinj i P) c = (let (R, S) = cfactor P c in (mkPinj i R, mkPinj i S))" | "cfactor (PX P i Q) c = (let (R1, S1) = cfactor P c; (R2, S2) = cfactor Q c in (mkPX R1 i R2, mkPX S1 i S2))" lemma (in cring) cfactor_correct: "in_carrier l \ Ipol l P = Ipol l (fst (cfactor P c)) \ \c\ \ Ipol l (snd (cfactor P c))" proof (induct P c arbitrary: l rule: cfactor.induct) case (1 c' c) show ?case by (simp add: of_int_mult [symmetric] of_int_add [symmetric] del: of_int_mult) next case (2 i P c) then show ?case by (simp add: mkPinj_ci split_beta) next case (3 P i Q c) from 3(1) 3(2) [OF refl prod.collapse] 3(3) show ?case by (simp add: mkPX_ci r_distr a_ac m_ac split_beta) qed fun mfactor :: "pol \ mon \ pol \ pol" where "mfactor P (Mc c) = (if c = 1 then (Pc 0, P) else cfactor P c)" | "mfactor (Pc d) M = (Pc d, Pc 0)" | "mfactor (Pinj i P) (Minj j M) = (if i = j then let (R, S) = mfactor P M in (mkPinj i R, mkPinj i S) else if i < j then let (R, S) = mfactor P (Minj (j - i) M) in (mkPinj i R, mkPinj i S) else (Pinj i P, Pc 0))" | "mfactor (Pinj i P) (MX j M) = (Pinj i P, Pc 0)" | "mfactor (PX P i Q) (Minj j M) = (if j = 0 then mfactor (PX P i Q) M else let (R1, S1) = mfactor P (Minj j M); (R2, S2) = mfactor Q (Minj_pred j M) in (mkPX R1 i R2, mkPX S1 i S2))" | "mfactor (PX P i Q) (MX j M) = (if i = j then let (R, S) = mfactor P (mkMinj 1 M) in (mkPX R i Q, S) else if i < j then let (R, S) = mfactor P (MX (j - i) M) in (mkPX R i Q, S) else let (R, S) = mfactor P (mkMinj 1 M) in (mkPX R i Q, mkPX S (i - j) (Pc 0)))" lemmas mfactor_induct = mfactor.induct [case_names Mc Pc_Minj Pc_MX Pinj_Minj Pinj_MX PX_Minj PX_MX] lemma (in cring) mfactor_correct: "in_carrier l \ Ipol l P = Ipol l (fst (mfactor P M)) \ Imon l M \ Ipol l (snd (mfactor P M))" proof (induct P M arbitrary: l rule: mfactor_induct) case (Mc P c) then show ?case by (simp add: cfactor_correct) next case (Pc_Minj d i M) then show ?case by simp next case (Pc_MX d i M) then show ?case by simp next case (Pinj_Minj i P j M) then show ?case by (simp add: mkPinj_ci split_beta) next case (Pinj_MX i P j M) then show ?case by simp next case (PX_Minj P i Q j M) from PX_Minj(1,2) PX_Minj(3) [OF _ refl prod.collapse] PX_Minj(4) show ?case by (simp add: mkPX_ci Minj_pred_correct [simplified] r_distr a_ac m_ac split_beta) next case (PX_MX P i Q j M) from \in_carrier l\ have eq1: "(Imon (drop (Suc 0) l) M \ head l [^] (j - i)) \ Ipol l (snd (mfactor P (MX (j - i) M))) \ head l [^] i = Imon (drop (Suc 0) l) M \ Ipol l (snd (mfactor P (MX (j - i) M))) \ (head l [^] (j - i) \ head l [^] i)" by (simp add: m_ac) from \in_carrier l\ have eq2: "(Imon (drop (Suc 0) l) M \ head l [^] j) \ (Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \ head l [^] (i - j)) = Imon (drop (Suc 0) l) M \ Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \ (head l [^] (i - j) \ head l [^] j)" by (simp add: m_ac) from PX_MX \in_carrier l\ show ?case by (simp add: mkPX_ci mkMinj_correct l_distr eq1 eq2 split_beta nat_pow_mult) (simp add: a_ac m_ac) qed primrec mon_of_pol :: "pol \ mon option" where "mon_of_pol (Pc c) = Some (Mc c)" | "mon_of_pol (Pinj i P) = (case mon_of_pol P of None \ None | Some M \ Some (mkMinj i M))" | "mon_of_pol (PX P i Q) = (if Q = Pc 0 then (case mon_of_pol P of None \ None | Some M \ Some (mkMX i M)) else None)" lemma (in cring) mon_of_pol_correct: assumes "in_carrier l" and "mon_of_pol P = Some M" shows "Ipol l P = Imon l M" using assms proof (induct P arbitrary: M l) case (PX P1 i P2) from PX(1,3,4) show ?case by (auto simp add: mkMinj_correct mkMX_correct split: if_split_asm option.split_asm) qed (auto simp add: mkMinj_correct split: option.split_asm) fun (in cring) Ipolex_polex_list :: "'a list \ (polex \ polex) list \ bool" where "Ipolex_polex_list l [] = True" | "Ipolex_polex_list l ((P, Q) # pps) = ((Ipolex l P = Ipolex l Q) \ Ipolex_polex_list l pps)" fun (in cring) Imon_pol_list :: "'a list \ (mon \ pol) list \ bool" where "Imon_pol_list l [] = True" | "Imon_pol_list l ((M, P) # mps) = ((Imon l M = Ipol l P) \ Imon_pol_list l mps)" fun mk_monpol_list :: "(polex \ polex) list \ (mon \ pol) list" where "mk_monpol_list [] = []" | "mk_monpol_list ((P, Q) # pps) = (case mon_of_pol (norm P) of None \ mk_monpol_list pps | Some M \ (M, norm Q) # mk_monpol_list pps)" lemma (in cring) mk_monpol_list_correct: "in_carrier l \ Ipolex_polex_list l pps \ Imon_pol_list l (mk_monpol_list pps)" by (induct pps rule: mk_monpol_list.induct) (auto split: option.split simp add: norm_ci [symmetric] mon_of_pol_correct [symmetric]) definition ponesubst :: "pol \ mon \ pol \ pol option" where "ponesubst P1 M P2 = (let (Q, R) = mfactor P1 M in (case R of Pc c \ if c = 0 then None else Some (add Q (mul P2 R)) | _ \ Some (add Q (mul P2 R))))" fun pnsubst1 :: "pol \ mon \ pol \ nat \ pol" where "pnsubst1 P1 M P2 n = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ if n = 0 then P3 else pnsubst1 P3 M P2 (n - 1))" lemma pnsubst1_0 [simp]: "pnsubst1 P1 M P2 0 = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ P3)" by (simp split: option.split) lemma pnsubst1_Suc [simp]: "pnsubst1 P1 M P2 (Suc n) = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ pnsubst1 P3 M P2 n)" by (simp split: option.split) declare pnsubst1.simps [simp del] definition pnsubst :: "pol \ mon \ pol \ nat \ pol option" where "pnsubst P1 M P2 n = (case ponesubst P1 M P2 of None \ None | Some P3 \ Some (pnsubst1 P3 M P2 n))" fun psubstl1 :: "pol \ (mon \ pol) list \ nat \ pol" where "psubstl1 P1 [] n = P1" | "psubstl1 P1 ((M, P2) # mps) n = psubstl1 (pnsubst1 P1 M P2 n) mps n" fun psubstl :: "pol \ (mon \ pol) list \ nat \ pol option" where "psubstl P1 [] n = None" | "psubstl P1 ((M, P2) # mps) n = (case pnsubst P1 M P2 n of None \ psubstl P1 mps n | Some P3 \ Some (psubstl1 P3 mps n))" fun pnsubstl :: "pol \ (mon \ pol) list \ nat \ nat \ pol" where "pnsubstl P1 mps m n = (case psubstl P1 mps n of None \ P1 | Some P3 \ if m = 0 then P3 else pnsubstl P3 mps (m - 1) n)" lemma pnsubstl_0 [simp]: "pnsubstl P1 mps 0 n = (case psubstl P1 mps n of None \ P1 | Some P3 \ P3)" by (simp split: option.split) lemma pnsubstl_Suc [simp]: "pnsubstl P1 mps (Suc m) n = (case psubstl P1 mps n of None \ P1 | Some P3 \ pnsubstl P3 mps m n)" by (simp split: option.split) declare pnsubstl.simps [simp del] lemma (in cring) ponesubst_correct: "in_carrier l \ ponesubst P1 M P2 = Some P3 \ Imon l M = Ipol l P2 \ Ipol l P1 = Ipol l P3" by (auto simp add: ponesubst_def split_beta mfactor_correct [of l P1 M] add_ci mul_ci split: pol.split_asm if_split_asm) lemma (in cring) pnsubst1_correct: "in_carrier l \ Imon l M = Ipol l P2 \ Ipol l (pnsubst1 P1 M P2 n) = Ipol l P1" by (induct n arbitrary: P1) (simp_all add: ponesubst_correct split: option.split) lemma (in cring) pnsubst_correct: "in_carrier l \ pnsubst P1 M P2 n = Some P3 \ Imon l M = Ipol l P2 \ Ipol l P1 = Ipol l P3" by (auto simp add: pnsubst_def pnsubst1_correct ponesubst_correct split: option.split_asm) lemma (in cring) psubstl1_correct: "in_carrier l \ Imon_pol_list l mps \ Ipol l (psubstl1 P1 mps n) = Ipol l P1" by (induct P1 mps n rule: psubstl1.induct) (simp_all add: pnsubst1_correct) lemma (in cring) psubstl_correct: "in_carrier l \ psubstl P1 mps n = Some P2 \ Imon_pol_list l mps \ Ipol l P1 = Ipol l P2" by (induct P1 mps n rule: psubstl.induct) (auto simp add: psubstl1_correct pnsubst_correct split: option.split_asm) lemma (in cring) pnsubstl_correct: "in_carrier l \ Imon_pol_list l mps \ Ipol l (pnsubstl P1 mps m n) = Ipol l P1" by (induct m arbitrary: P1) (simp_all add: psubstl_correct split: option.split) lemma (in cring) norm_subst_correct: "in_carrier l \ Ipolex_polex_list l pps \ Ipolex l P = Ipol l (pnsubstl (norm P) (mk_monpol_list pps) m n)" by (simp add: pnsubstl_correct mk_monpol_list_correct norm_ci) lemma in_carrier_trivial: "cring_class.in_carrier l" by (induct l) (simp_all add: cring_class.in_carrier_def carrier_class) ML \ val term_of_nat = HOLogic.mk_number \<^Type>\nat\ o @{code integer_of_nat}; val term_of_int = HOLogic.mk_number \<^Type>\int\ o @{code integer_of_int}; fun term_of_pol (@{code Pc} k) = \<^Const>\Pc\ $ term_of_int k | term_of_pol (@{code Pinj} (n, p)) = \<^Const>\Pinj\ $ term_of_nat n $ term_of_pol p | term_of_pol (@{code PX} (p, n, q)) = \<^Const>\PX\ $ term_of_pol p $ term_of_nat n $ term_of_pol q; local fun pol (ctxt, ct, t) = \<^instantiate>\x = ct and y = \Thm.cterm_of ctxt t\ in cterm \x \ y\ for x y :: pol\; val (_, raw_pol_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\pnsubstl\, pol))); fun pol_oracle ctxt ct t = raw_pol_oracle (ctxt, ct, t); in val cv = @{computation_conv pol terms: pnsubstl mk_monpol_list norm nat_of_integer Code_Target_Nat.natural "0::nat" "1::nat" "2::nat" "3::nat" "0::int" "1::int" "2::int" "3::int" "-1::int" datatypes: polex "(polex * polex) list" int integer num} (fn ctxt => fn p => fn ct => pol_oracle ctxt ct (term_of_pol p)) end \ ML \ signature RING_TAC = sig structure Ring_Simps: sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end val insert_rules: ((term * 'a) * (term * 'a) -> bool) -> (term * 'a) -> (term * 'a) Net.net -> (term * 'a) Net.net val eq_ring_simps: (term * (thm list * thm list * thm list * thm list * thm * thm)) * (term * (thm list * thm list * thm list * thm list * thm * thm)) -> bool val ring_tac: bool -> thm list -> Proof.context -> int -> tactic val get_matching_rules: Proof.context -> (term * 'a) Net.net -> term -> 'a option val norm: thm -> thm val mk_in_carrier: Proof.context -> term -> thm list -> (string * typ) list -> thm val mk_ring: typ -> term end structure Ring_Tac : RING_TAC = struct fun eq_ring_simps ((t, (ths1, ths2, ths3, ths4, th5, th)), (t', (ths1', ths2', ths3', ths4', th5', th'))) = t aconv t' andalso eq_list Thm.eq_thm (ths1, ths1') andalso eq_list Thm.eq_thm (ths2, ths2') andalso eq_list Thm.eq_thm (ths3, ths3') andalso eq_list Thm.eq_thm (ths4, ths4') andalso Thm.eq_thm (th5, th5') andalso Thm.eq_thm (th, th'); structure Ring_Simps = Generic_Data (struct type T = (term * (thm list * thm list * thm list * thm list * thm * thm)) Net.net val empty = Net.empty val merge = Net.merge eq_ring_simps end); fun get_matching_rules ctxt net t = get_first (fn (p, x) => if Pattern.matches (Proof_Context.theory_of ctxt) (p, t) then SOME x else NONE) (Net.match_term net t); fun insert_rules eq (t, x) = Net.insert_term eq (t, (t, x)); fun norm thm = thm COMP_INCR asm_rl; fun get_ring_simps ctxt optcT t = (case get_matching_rules ctxt (Ring_Simps.get (Context.Proof ctxt)) t of SOME (ths1, ths2, ths3, ths4, th5, th) => let val tr = Thm.transfer' ctxt #> (case optcT of NONE => I | SOME cT => inst [cT] [] #> norm) in (map tr ths1, map tr ths2, map tr ths3, map tr ths4, tr th5, tr th) end | NONE => error "get_ring_simps: lookup failed"); fun ring_struct \<^Const_>\Ring.ring.add _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.a_minus _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Group.monoid.mult _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.a_inv _ _ for R _\ = SOME R | ring_struct \<^Const_>\Group.pow _ _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.ring.zero _ _ for R\ = SOME R | ring_struct \<^Const_>\Group.monoid.one _ _ for R\ = SOME R | ring_struct \<^Const_>\Algebra_Aux.of_integer _ _ for R _\ = SOME R | ring_struct _ = NONE; fun reif_polex vs \<^Const_>\Ring.ring.add _ _ for _ a b\ = \<^Const>\Add for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Ring.a_minus _ _ for _ a b\ = \<^Const>\Sub for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Group.monoid.mult _ _ for _ a b\ = \<^Const>\Mul for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Ring.a_inv _ _ for _ a\ = \<^Const>\Neg for \reif_polex vs a\\ | reif_polex vs \<^Const_>\Group.pow _ _ _ for _ a n\ = \<^Const>\Pow for \reif_polex vs a\ n\ | reif_polex vs (Free x) = \<^Const>\Var for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_polex _ \<^Const_>\Ring.ring.zero _ _ for _\ = \<^term>\Const 0\ | reif_polex _ \<^Const_>\Group.monoid.one _ _ for _\ = \<^term>\Const 1\ | reif_polex _ \<^Const_>\Algebra_Aux.of_integer _ _ for _ n\ = \<^Const>\Const for n\ | reif_polex _ _ = error "reif_polex: bad expression"; fun reif_polex' vs \<^Const_>\plus _ for a b\ = \<^Const>\Add for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\minus _ for a b\ = \<^Const>\Sub for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\times _ for a b\ = \<^Const>\Mul for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\uminus _ for a\ = \<^Const>\Neg for \reif_polex' vs a\\ | reif_polex' vs \<^Const_>\power _ for a n\ = \<^Const>\Pow for \reif_polex' vs a\ n\ | reif_polex' vs (Free x) = \<^Const>\Var for \HOLogic.mk_number \<^Type>\nat\ (find_index (equal x) vs)\\ | reif_polex' _ \<^Const_>\numeral _ for b\ = \<^Const>\Const for \<^Const>\numeral \<^Type>\int\ for b\\ | reif_polex' _ \<^Const_>\zero_class.zero _\ = \<^term>\Const 0\ | reif_polex' _ \<^Const_>\one_class.one _\ = \<^term>\Const 1\ | reif_polex' _ t = error "reif_polex: bad expression"; fun head_conv (_, _, _, _, head_simp, _) ys = (case strip_app ys of (\<^const_name>\Cons\, [y, xs]) => inst [] [y, xs] head_simp); fun Ipol_conv (rls as ([Ipol_simps_1, Ipol_simps_2, Ipol_simps_3, Ipol_simps_4, Ipol_simps_5, Ipol_simps_6, Ipol_simps_7], _, _, _, _, _)) = let val a = type_of_eqn Ipol_simps_1; val drop_conv_a = drop_conv a; fun conv l p = (case strip_app p of (\<^const_name>\Pc\, [c]) => (case strip_numeral c of (\<^const_name>\zero_class.zero\, _) => inst [] [l] Ipol_simps_4 | (\<^const_name>\one_class.one\, _) => inst [] [l] Ipol_simps_5 | (\<^const_name>\numeral\, [m]) => inst [] [l, m] Ipol_simps_6 | (\<^const_name>\uminus\, [m]) => inst [] [l, m] Ipol_simps_7 | _ => inst [] [l, c] Ipol_simps_1) | (\<^const_name>\Pinj\, [i, P]) => transitive' (inst [] [l, i, P] Ipol_simps_2) (cong2' conv (args2 drop_conv_a) Thm.reflexive) | (\<^const_name>\PX\, [P, x, Q]) => transitive' (inst [] [l, P, x, Q] Ipol_simps_3) (cong2 (cong2 (args2 conv) (cong2 (args1 (head_conv rls)) Thm.reflexive)) (cong2' conv (args2 drop_conv_a) Thm.reflexive))) in conv end; fun Ipolex_conv (rls as (_, [Ipolex_Var, Ipolex_Const, Ipolex_Add, Ipolex_Sub, Ipolex_Mul, Ipolex_Pow, Ipolex_Neg, Ipolex_Const_0, Ipolex_Const_1, Ipolex_Const_numeral], _, _, _, _)) = let val a = type_of_eqn Ipolex_Var; val drop_conv_a = drop_conv a; fun conv l r = (case strip_app r of (\<^const_name>\Var\, [n]) => transitive' (inst [] [l, n] Ipolex_Var) (cong1' (head_conv rls) (args2 drop_conv_a)) | (\<^const_name>\Const\, [i]) => (case strip_app i of (\<^const_name>\zero_class.zero\, _) => inst [] [l] Ipolex_Const_0 | (\<^const_name>\one_class.one\, _) => inst [] [l] Ipolex_Const_1 | (\<^const_name>\numeral\, [m]) => inst [] [l, m] Ipolex_Const_numeral | _ => inst [] [l, i] Ipolex_Const) | (\<^const_name>\Add\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Add) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Sub\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Sub) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Mul\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Mul) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Pow\, [P, n]) => transitive' (inst [] [l, P, n] Ipolex_Pow) (cong2 (args2 conv) Thm.reflexive) | (\<^const_name>\Neg\, [P]) => transitive' (inst [] [l, P] Ipolex_Neg) (cong1 (args2 conv))) in conv end; fun Ipolex_polex_list_conv (rls as (_, _, [Ipolex_polex_list_Nil, Ipolex_polex_list_Cons], _, _, _)) l pps = (case strip_app pps of (\<^const_name>\Nil\, []) => inst [] [l] Ipolex_polex_list_Nil | (\<^const_name>\Cons\, [p, pps']) => (case strip_app p of (\<^const_name>\Pair\, [P, Q]) => transitive' (inst [] [l, P, Q, pps'] Ipolex_polex_list_Cons) (cong2 (cong2 (args2 (Ipolex_conv rls)) (args2 (Ipolex_conv rls))) (args2 (Ipolex_polex_list_conv rls))))); fun dest_conj th = let val th1 = th RS @{thm conjunct1}; val th2 = th RS @{thm conjunct2} in dest_conj th1 @ dest_conj th2 end handle THM _ => [th]; fun mk_in_carrier ctxt R prems xs = let val (_, _, _, [in_carrier_Nil, in_carrier_Cons], _, _) = get_ring_simps ctxt NONE R; val props = map fst (Facts.props (Proof_Context.facts_of ctxt)) @ maps dest_conj prems; val ths = map (fn p as (x, _) => (case find_first ((fn \<^Const_>\Trueprop for \<^Const_>\Set.member _ for \Free (y, _)\ \<^Const_>\carrier _ _ for S\\\ => x = y andalso R aconv S | _ => false) o Thm.prop_of) props of SOME th => th | NONE => error ("Variable " ^ Syntax.string_of_term ctxt (Free p) ^ " not in carrier"))) xs in fold_rev (fn th1 => fn th2 => [th1, th2] MRS in_carrier_Cons) ths in_carrier_Nil end; fun mk_ring T = \<^Const>\cring_class_ops T\; val iterations = \<^cterm>\1000::nat\; val Trueprop_cong = Thm.combination (Thm.reflexive \<^cterm>\Trueprop\); fun commutative_ring_conv ctxt prems eqs ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; val eqs' = map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) eqs; val xs = filter (equal T o snd) (rev (fold Term.add_frees (map fst eqs' @ map snd eqs' @ [Thm.term_of ct]) [])); val (R, optcT, prem', reif) = (case ring_struct (Thm.term_of ct) of SOME R => (R, NONE, mk_in_carrier ctxt R prems xs, reif_polex xs) | NONE => (mk_ring T, SOME cT, @{thm in_carrier_trivial}, reif_polex' xs)); val rls as (_, _, _, _, _, norm_subst_correct) = get_ring_simps ctxt optcT R; val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs)); val ceqs = Thm.cterm_of ctxt (HOLogic.mk_list \<^typ>\polex \ polex\ (map (HOLogic.mk_prod o apply2 reif) eqs')); val cp = Thm.cterm_of ctxt (reif (Thm.term_of ct)); val prem = Thm.equal_elim (Trueprop_cong (Thm.symmetric (Ipolex_polex_list_conv rls cxs ceqs))) (fold_rev (fn th1 => fn th2 => [th1, th2] MRS @{thm conjI}) eqs @{thm TrueI}); in Thm.transitive (Thm.symmetric (Ipolex_conv rls cxs cp)) (transitive' ([prem', prem] MRS inst [] [cxs, ceqs, cp, iterations, iterations] norm_subst_correct) (cong2' (Ipol_conv rls) Thm.reflexive (cv ctxt))) end; fun ring_tac in_prems thms ctxt = tactic_of_conv (fn ct => (if in_prems then Conv.prems_conv else Conv.concl_conv) (Logic.count_prems (Thm.term_of ct)) (Conv.arg_conv (Conv.binop_conv (commutative_ring_conv ctxt [] thms))) ct) THEN' TRY o (assume_tac ctxt ORELSE' resolve_tac ctxt [@{thm refl}]); end \ context cring begin local_setup \ -Local_Theory.declaration {syntax = false, pervasive = false} +Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Ring_Tac.Ring_Simps.map (Ring_Tac.insert_rules Ring_Tac.eq_ring_simps (Morphism.term phi \<^term>\R\, (Morphism.fact phi @{thms Ipol.simps [meta] Ipol_Pc [meta]}, Morphism.fact phi @{thms Ipolex.simps [meta] Ipolex_Const [meta]}, Morphism.fact phi @{thms Ipolex_polex_list.simps [meta]}, Morphism.fact phi @{thms in_carrier_Nil in_carrier_Cons}, singleton (Morphism.fact phi) @{thm head.simps(2) [meta]}, singleton (Morphism.fact phi) @{thm norm_subst_correct [meta]})))) \ end method_setup ring = \ Scan.lift (Args.mode "prems") -- Attrib.thms >> (SIMPLE_METHOD' oo uncurry Ring_Tac.ring_tac) \ "simplify equations involving rings" end diff --git a/src/HOL/Decision_Procs/Reflective_Field.thy b/src/HOL/Decision_Procs/Reflective_Field.thy --- a/src/HOL/Decision_Procs/Reflective_Field.thy +++ b/src/HOL/Decision_Procs/Reflective_Field.thy @@ -1,930 +1,930 @@ (* Title: HOL/Decision_Procs/Reflective_Field.thy Author: Stefan Berghofer Reducing equalities in fields to equalities in rings. *) theory Reflective_Field imports Commutative_Ring begin datatype fexpr = FCnst int | FVar nat | FAdd fexpr fexpr | FSub fexpr fexpr | FMul fexpr fexpr | FNeg fexpr | FDiv fexpr fexpr | FPow fexpr nat fun (in field) nth_el :: "'a list \ nat \ 'a" where "nth_el [] n = \" | "nth_el (x # xs) 0 = x" | "nth_el (x # xs) (Suc n) = nth_el xs n" lemma (in field) nth_el_Cons: "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))" by (cases n) simp_all lemma (in field) nth_el_closed [simp]: "in_carrier xs \ nth_el xs n \ carrier R" by (induct xs n rule: nth_el.induct) (simp_all add: in_carrier_def) primrec (in field) feval :: "'a list \ fexpr \ 'a" where "feval xs (FCnst c) = \c\" | "feval xs (FVar n) = nth_el xs n" | "feval xs (FAdd a b) = feval xs a \ feval xs b" | "feval xs (FSub a b) = feval xs a \ feval xs b" | "feval xs (FMul a b) = feval xs a \ feval xs b" | "feval xs (FNeg a) = \ feval xs a" | "feval xs (FDiv a b) = feval xs a \ feval xs b" | "feval xs (FPow a n) = feval xs a [^] n" lemma (in field) feval_Cnst: "feval xs (FCnst 0) = \" "feval xs (FCnst 1) = \" "feval xs (FCnst (numeral n)) = \numeral n\" by simp_all datatype pexpr = PExpr1 pexpr1 | PExpr2 pexpr2 and pexpr1 = PCnst int | PVar nat | PAdd pexpr pexpr | PSub pexpr pexpr | PNeg pexpr and pexpr2 = PMul pexpr pexpr | PPow pexpr nat lemma pexpr_cases [case_names PCnst PVar PAdd PSub PNeg PMul PPow]: assumes "\c. e = PExpr1 (PCnst c) \ P" "\n. e = PExpr1 (PVar n) \ P" "\e1 e2. e = PExpr1 (PAdd e1 e2) \ P" "\e1 e2. e = PExpr1 (PSub e1 e2) \ P" "\e'. e = PExpr1 (PNeg e') \ P" "\e1 e2. e = PExpr2 (PMul e1 e2) \ P" "\e' n. e = PExpr2 (PPow e' n) \ P" shows P proof (cases e) case (PExpr1 e') then show ?thesis apply (cases e') apply simp_all apply (erule assms)+ done next case (PExpr2 e') then show ?thesis apply (cases e') apply simp_all apply (erule assms)+ done qed lemmas pexpr_cases2 = pexpr_cases [case_product pexpr_cases] fun (in field) peval :: "'a list \ pexpr \ 'a" where "peval xs (PExpr1 (PCnst c)) = \c\" | "peval xs (PExpr1 (PVar n)) = nth_el xs n" | "peval xs (PExpr1 (PAdd a b)) = peval xs a \ peval xs b" | "peval xs (PExpr1 (PSub a b)) = peval xs a \ peval xs b" | "peval xs (PExpr1 (PNeg a)) = \ peval xs a" | "peval xs (PExpr2 (PMul a b)) = peval xs a \ peval xs b" | "peval xs (PExpr2 (PPow a n)) = peval xs a [^] n" lemma (in field) peval_Cnst: "peval xs (PExpr1 (PCnst 0)) = \" "peval xs (PExpr1 (PCnst 1)) = \" "peval xs (PExpr1 (PCnst (numeral n))) = \numeral n\" "peval xs (PExpr1 (PCnst (- numeral n))) = \ \numeral n\" by simp_all lemma (in field) peval_closed [simp]: "in_carrier xs \ peval xs e \ carrier R" "in_carrier xs \ peval xs (PExpr1 e1) \ carrier R" "in_carrier xs \ peval xs (PExpr2 e2) \ carrier R" by (induct e and e1 and e2) simp_all definition npepow :: "pexpr \ nat \ pexpr" where "npepow e n = (if n = 0 then PExpr1 (PCnst 1) else if n = 1 then e else (case e of PExpr1 (PCnst c) \ PExpr1 (PCnst (c ^ n)) | _ \ PExpr2 (PPow e n)))" lemma (in field) npepow_correct: "in_carrier xs \ peval xs (npepow e n) = peval xs (PExpr2 (PPow e n))" by (cases e rule: pexpr_cases) (simp_all add: npepow_def) fun npemul :: "pexpr \ pexpr \ pexpr" where "npemul x y = (case x of PExpr1 (PCnst c) \ if c = 0 then x else if c = 1 then y else (case y of PExpr1 (PCnst d) \ PExpr1 (PCnst (c * d)) | _ \ PExpr2 (PMul x y)) | PExpr2 (PPow e1 n) \ (case y of PExpr2 (PPow e2 m) \ if n = m then npepow (npemul e1 e2) n else PExpr2 (PMul x y) | PExpr1 (PCnst d) \ if d = 0 then y else if d = 1 then x else PExpr2 (PMul x y) | _ \ PExpr2 (PMul x y)) | _ \ (case y of PExpr1 (PCnst d) \ if d = 0 then y else if d = 1 then x else PExpr2 (PMul x y) | _ \ PExpr2 (PMul x y)))" lemma (in field) npemul_correct: "in_carrier xs \ peval xs (npemul e1 e2) = peval xs (PExpr2 (PMul e1 e2))" proof (induct e1 e2 rule: npemul.induct) case (1 x y) then show ?case proof (cases x y rule: pexpr_cases2) case (PPow_PPow e n e' m) then show ?thesis by (simp del: npemul.simps add: 1 npepow_correct nat_pow_distrib npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"]) qed simp_all qed declare npemul.simps [simp del] definition npeadd :: "pexpr \ pexpr \ pexpr" where "npeadd x y = (case x of PExpr1 (PCnst c) \ if c = 0 then y else (case y of PExpr1 (PCnst d) \ PExpr1 (PCnst (c + d)) | _ \ PExpr1 (PAdd x y)) | _ \ (case y of PExpr1 (PCnst d) \ if d = 0 then x else PExpr1 (PAdd x y) | _ \ PExpr1 (PAdd x y)))" lemma (in field) npeadd_correct: "in_carrier xs \ peval xs (npeadd e1 e2) = peval xs (PExpr1 (PAdd e1 e2))" by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npeadd_def) definition npesub :: "pexpr \ pexpr \ pexpr" where "npesub x y = (case y of PExpr1 (PCnst d) \ if d = 0 then x else (case x of PExpr1 (PCnst c) \ PExpr1 (PCnst (c - d)) | _ \ PExpr1 (PSub x y)) | _ \ (case x of PExpr1 (PCnst c) \ if c = 0 then PExpr1 (PNeg y) else PExpr1 (PSub x y) | _ \ PExpr1 (PSub x y)))" lemma (in field) npesub_correct: "in_carrier xs \ peval xs (npesub e1 e2) = peval xs (PExpr1 (PSub e1 e2))" by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npesub_def) definition npeneg :: "pexpr \ pexpr" where "npeneg e = (case e of PExpr1 (PCnst c) \ PExpr1 (PCnst (- c)) | _ \ PExpr1 (PNeg e))" lemma (in field) npeneg_correct: "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))" by (cases e rule: pexpr_cases) (simp_all add: npeneg_def) lemma option_pair_cases [case_names None Some]: obtains (None) "x = None" | (Some) p q where "x = Some (p, q)" proof (cases x) case None then show ?thesis by (rule that) next case (Some r) then show ?thesis by (cases r, simp) (rule that) qed fun isin :: "pexpr \ nat \ pexpr \ nat \ (nat \ pexpr) option" where "isin e n (PExpr2 (PMul e1 e2)) m = (case isin e n e1 m of Some (k, e3) \ if k = 0 then Some (0, npemul e3 (npepow e2 m)) else (case isin e k e2 m of Some (l, e4) \ Some (l, npemul e3 e4) | None \ Some (k, npemul e3 (npepow e2 m))) | None \ (case isin e n e2 m of Some (k, e3) \ Some (k, npemul (npepow e1 m) e3) | None \ None))" | "isin e n (PExpr2 (PPow e' k)) m = (if k = 0 then None else isin e n e' (k * m))" | "isin (PExpr1 e) n (PExpr1 e') m = (if e = e' then if n \ m then Some (n - m, PExpr1 (PCnst 1)) else Some (0, npepow (PExpr1 e) (m - n)) else None)" | "isin (PExpr2 e) n (PExpr1 e') m = None" lemma (in field) isin_correct: assumes "in_carrier xs" and "isin e n e' m = Some (p, e'')" shows "peval xs (PExpr2 (PPow e' m)) = peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))" and "p \ n" using assms by (induct e n e' m arbitrary: p e'' rule: isin.induct) (force simp add: nat_pow_distrib nat_pow_pow nat_pow_mult m_ac npemul_correct npepow_correct split: option.split_asm prod.split_asm if_split_asm)+ lemma (in field) isin_correct': "in_carrier xs \ isin e n e' 1 = Some (p, e'') \ peval xs e' = peval xs e [^] (n - p) \ peval xs e''" "in_carrier xs \ isin e n e' 1 = Some (p, e'') \ p \ n" using isin_correct [where m = 1] by simp_all fun split_aux :: "pexpr \ nat \ pexpr \ pexpr \ pexpr \ pexpr" where "split_aux (PExpr2 (PMul e1 e2)) n e = (let (left1, common1, right1) = split_aux e1 n e; (left2, common2, right2) = split_aux e2 n right1 in (npemul left1 left2, npemul common1 common2, right2))" | "split_aux (PExpr2 (PPow e' m)) n e = (if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e) else split_aux e' (m * n) e)" | "split_aux (PExpr1 e') n e = (case isin (PExpr1 e') n e 1 of Some (m, e'') \ (if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'') else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e'')) | None \ (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))" hide_const Left Right (* FIXME !? *) abbreviation Left :: "pexpr \ pexpr \ pexpr" where "Left e1 e2 \ fst (split_aux e1 (Suc 0) e2)" abbreviation Common :: "pexpr \ pexpr \ pexpr" where "Common e1 e2 \ fst (snd (split_aux e1 (Suc 0) e2))" abbreviation Right :: "pexpr \ pexpr \ pexpr" where "Right e1 e2 \ snd (snd (split_aux e1 (Suc 0) e2))" lemma split_aux_induct [case_names 1 2 3]: assumes I1: "\e1 e2 n e. P e1 n e \ P e2 n (snd (snd (split_aux e1 n e))) \ P (PExpr2 (PMul e1 e2)) n e" and I2: "\e' m n e. (m \ 0 \ P e' (m * n) e) \ P (PExpr2 (PPow e' m)) n e" and I3: "\e' n e. P (PExpr1 e') n e" shows "P x y z" proof (induct x y z rule: split_aux.induct) case 1 from 1(1) 1(2) [OF refl prod.collapse prod.collapse] show ?case by (rule I1) next case 2 then show ?case by (rule I2) next case 3 then show ?case by (rule I3) qed lemma (in field) split_aux_correct: "in_carrier xs \ peval xs (PExpr2 (PPow e\<^sub>1 n)) = peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))" "in_carrier xs \ peval xs e\<^sub>2 = peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))" by (induct e\<^sub>1 n e\<^sub>2 rule: split_aux_induct) (auto simp add: split_beta nat_pow_distrib nat_pow_pow nat_pow_mult m_ac npemul_correct npepow_correct isin_correct' split: option.split) lemma (in field) split_aux_correct': "in_carrier xs \ peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \ peval xs (Common e\<^sub>1 e\<^sub>2)" "in_carrier xs \ peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \ peval xs (Common e\<^sub>1 e\<^sub>2)" using split_aux_correct [where n = 1] by simp_all fun fnorm :: "fexpr \ pexpr \ pexpr \ pexpr list" where "fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])" | "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])" | "fnorm (FAdd e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left, common, right) = split_aux xd 1 yd in (npeadd (npemul xn right) (npemul yn left), npemul left (npemul right common), List.union xc yc))" | "fnorm (FSub e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left, common, right) = split_aux xd 1 yd in (npesub (npemul xn right) (npemul yn left), npemul left (npemul right common), List.union xc yc))" | "fnorm (FMul e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left1, common1, right1) = split_aux xn 1 yd; (left2, common2, right2) = split_aux yn 1 xd in (npemul left1 left2, npemul right2 right1, List.union xc yc))" | "fnorm (FNeg e) = (let (n, d, c) = fnorm e in (npeneg n, d, c))" | "fnorm (FDiv e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left1, common1, right1) = split_aux xn 1 yn; (left2, common2, right2) = split_aux xd 1 yd in (npemul left1 right2, npemul left2 right1, List.insert yn (List.union xc yc)))" | "fnorm (FPow e m) = (let (n, d, c) = fnorm e in (npepow n m, npepow d m, c))" abbreviation Num :: "fexpr \ pexpr" where "Num e \ fst (fnorm e)" abbreviation Denom :: "fexpr \ pexpr" where "Denom e \ fst (snd (fnorm e))" abbreviation Cond :: "fexpr \ pexpr list" where "Cond e \ snd (snd (fnorm e))" primrec (in field) nonzero :: "'a list \ pexpr list \ bool" where "nonzero xs [] \ True" | "nonzero xs (p # ps) \ peval xs p \ \ \ nonzero xs ps" lemma (in field) nonzero_singleton: "nonzero xs [p] = (peval xs p \ \)" by simp lemma (in field) nonzero_append: "nonzero xs (ps @ qs) = (nonzero xs ps \ nonzero xs qs)" by (induct ps) simp_all lemma (in field) nonzero_idempotent: "p \ set ps \ (peval xs p \ \ \ nonzero xs ps) = nonzero xs ps" by (induct ps) auto lemma (in field) nonzero_insert: "nonzero xs (List.insert p ps) = (peval xs p \ \ \ nonzero xs ps)" by (simp add: List.insert_def nonzero_idempotent) lemma (in field) nonzero_union: "nonzero xs (List.union ps qs) = (nonzero xs ps \ nonzero xs qs)" by (induct ps rule: rev_induct) (auto simp add: List.union_def nonzero_insert nonzero_append) lemma (in field) fnorm_correct: assumes "in_carrier xs" and "nonzero xs (Cond e)" shows "feval xs e = peval xs (Num e) \ peval xs (Denom e)" and "peval xs (Denom e) \ \" using assms proof (induct e) case (FCnst c) { case 1 show ?case by simp next case 2 show ?case by simp } next case (FVar n) { case 1 then show ?case by simp next case 2 show ?case by simp } next case (FAdd e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left = "peval xs (Left (Denom e1) (Denom e2))" let ?common = "peval xs (Common (Denom e1) (Denom e2))" let ?right = "peval xs (Right (Denom e1) (Denom e2))" from 1 FAdd have "feval xs (FAdd e1 e2) = (?common \ (peval xs (Num e1) \ ?right \ peval xs (Num e2) \ ?left)) \ (?common \ (?left \ (?right \ ?common)))" by (simp add: split_beta split nonzero_union add_frac_eq r_distr m_ac) also from 1 FAdd have "\ = peval xs (Num (FAdd e1 e2)) \ peval xs (Denom (FAdd e1 e2))" by (simp add: split_beta split nonzero_union npeadd_correct npemul_correct integral_iff) finally show ?case . next case 2 with FAdd show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FSub e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left = "peval xs (Left (Denom e1) (Denom e2))" let ?common = "peval xs (Common (Denom e1) (Denom e2))" let ?right = "peval xs (Right (Denom e1) (Denom e2))" from 1 FSub have "feval xs (FSub e1 e2) = (?common \ (peval xs (Num e1) \ ?right \ peval xs (Num e2) \ ?left)) \ (?common \ (?left \ (?right \ ?common)))" by (simp add: split_beta split nonzero_union diff_frac_eq r_diff_distr m_ac) also from 1 FSub have "\ = peval xs (Num (FSub e1 e2)) \ peval xs (Denom (FSub e1 e2))" by (simp add: split_beta split nonzero_union npesub_correct npemul_correct integral_iff) finally show ?case . next case 2 with FSub show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FMul e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Denom e2"] split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e2" and e\<^sub>2 = "Denom e1"] { case 1 let ?left\<^sub>1 = "peval xs (Left (Num e1) (Denom e2))" let ?common\<^sub>1 = "peval xs (Common (Num e1) (Denom e2))" let ?right\<^sub>1 = "peval xs (Right (Num e1) (Denom e2))" let ?left\<^sub>2 = "peval xs (Left (Num e2) (Denom e1))" let ?common\<^sub>2 = "peval xs (Common (Num e2) (Denom e1))" let ?right\<^sub>2 = "peval xs (Right (Num e2) (Denom e1))" from 1 FMul have "feval xs (FMul e1 e2) = ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>1 \ ?left\<^sub>2)) \ ((?common\<^sub>1 \ ?common\<^sub>2) \ (?right\<^sub>2 \ ?right\<^sub>1))" by (simp add: split_beta split nonzero_union nonzero_divide_divide_eq_left m_ac) also from 1 FMul have "\ = peval xs (Num (FMul e1 e2)) \ peval xs (Denom (FMul e1 e2))" by (simp add: split_beta split nonzero_union npemul_correct integral_iff) finally show ?case . next case 2 with FMul show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FNeg e) { case 1 with FNeg show ?case by (simp add: split_beta npeneg_correct) next case 2 with FNeg show ?case by (simp add: split_beta) } next case (FDiv e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Num e2"] split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left\<^sub>1 = "peval xs (Left (Num e1) (Num e2))" let ?common\<^sub>1 = "peval xs (Common (Num e1) (Num e2))" let ?right\<^sub>1 = "peval xs (Right (Num e1) (Num e2))" let ?left\<^sub>2 = "peval xs (Left (Denom e1) (Denom e2))" let ?common\<^sub>2 = "peval xs (Common (Denom e1) (Denom e2))" let ?right\<^sub>2 = "peval xs (Right (Denom e1) (Denom e2))" from 1 FDiv have "feval xs (FDiv e1 e2) = ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>1 \ ?right\<^sub>2)) \ ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>2 \ ?right\<^sub>1))" by (simp add: split_beta split nonzero_union nonzero_insert nonzero_divide_divide_eq m_ac) also from 1 FDiv have "\ = peval xs (Num (FDiv e1 e2)) \ peval xs (Denom (FDiv e1 e2))" by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff) finally show ?case . next case 2 with FDiv show ?case by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff) } next case (FPow e n) { case 1 with FPow show ?case by (simp add: split_beta nonzero_power_divide npepow_correct) next case 2 with FPow show ?case by (simp add: split_beta npepow_correct) } qed lemma (in field) feval_eq0: assumes "in_carrier xs" and "fnorm e = (n, d, c)" and "nonzero xs c" and "peval xs n = \" shows "feval xs e = \" using assms fnorm_correct [of xs e] by simp lemma (in field) fexpr_in_carrier: assumes "in_carrier xs" and "nonzero xs (Cond e)" shows "feval xs e \ carrier R" using assms proof (induct e) case (FDiv e1 e2) then have "feval xs e1 \ carrier R" "feval xs e2 \ carrier R" "peval xs (Num e2) \ \" "nonzero xs (Cond e2)" by (simp_all add: nonzero_union nonzero_insert split: prod.split_asm) from \in_carrier xs\ \nonzero xs (Cond e2)\ have "feval xs e2 = peval xs (Num e2) \ peval xs (Denom e2)" by (rule fnorm_correct) moreover from \in_carrier xs\ \nonzero xs (Cond e2)\ have "peval xs (Denom e2) \ \" by (rule fnorm_correct) ultimately have "feval xs e2 \ \" using \peval xs (Num e2) \ \\ \in_carrier xs\ by (simp add: divide_eq_0_iff) with \feval xs e1 \ carrier R\ \feval xs e2 \ carrier R\ show ?case by simp qed (simp_all add: nonzero_union split: prod.split_asm) lemma (in field) feval_eq: assumes "in_carrier xs" and "fnorm (FSub e e') = (n, d, c)" and "nonzero xs c" shows "(feval xs e = feval xs e') = (peval xs n = \)" proof - from assms have "nonzero xs (Cond e)" "nonzero xs (Cond e')" by (auto simp add: nonzero_union split: prod.split_asm) with assms fnorm_correct [of xs "FSub e e'"] have "feval xs e \ feval xs e' = peval xs n \ peval xs d" "peval xs d \ \" by simp_all show ?thesis proof assume "feval xs e = feval xs e'" with \feval xs e \ feval xs e' = peval xs n \ peval xs d\ \in_carrier xs\ \nonzero xs (Cond e')\ have "peval xs n \ peval xs d = \" by (simp add: fexpr_in_carrier minus_eq r_neg) with \peval xs d \ \\ \in_carrier xs\ show "peval xs n = \" by (simp add: divide_eq_0_iff) next assume "peval xs n = \" with \feval xs e \ feval xs e' = peval xs n \ peval xs d\ \peval xs d \ \\ \nonzero xs (Cond e)\ \nonzero xs (Cond e')\ \in_carrier xs\ show "feval xs e = feval xs e'" by (simp add: eq_diff0 fexpr_in_carrier) qed qed ML \ val term_of_nat = HOLogic.mk_number \<^Type>\nat\ o @{code integer_of_nat}; val term_of_int = HOLogic.mk_number \<^Type>\int\ o @{code integer_of_int}; fun term_of_pexpr (@{code PExpr1} x) = \<^Const>\PExpr1\ $ term_of_pexpr1 x | term_of_pexpr (@{code PExpr2} x) = \<^Const>\PExpr2\ $ term_of_pexpr2 x and term_of_pexpr1 (@{code PCnst} k) = \<^Const>\PCnst\ $ term_of_int k | term_of_pexpr1 (@{code PVar} n) = \<^Const>\PVar\ $ term_of_nat n | term_of_pexpr1 (@{code PAdd} (x, y)) = \<^Const>\PAdd\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr1 (@{code PSub} (x, y)) = \<^Const>\PSub\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr1 (@{code PNeg} x) = \<^Const>\PNeg\ $ term_of_pexpr x and term_of_pexpr2 (@{code PMul} (x, y)) = \<^Const>\PMul\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr2 (@{code PPow} (x, n)) = \<^Const>\PPow\ $ term_of_pexpr x $ term_of_nat n fun term_of_result (x, (y, zs)) = HOLogic.mk_prod (term_of_pexpr x, HOLogic.mk_prod (term_of_pexpr y, HOLogic.mk_list \<^Type>\pexpr\ (map term_of_pexpr zs))); local fun fnorm (ctxt, ct, t) = \<^instantiate>\x = ct and y = \Thm.cterm_of ctxt t\ in cterm \x \ y\ for x y :: \pexpr \ pexpr \ pexpr list\\; val (_, raw_fnorm_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\fnorm\, fnorm))); fun fnorm_oracle ctxt ct t = raw_fnorm_oracle (ctxt, ct, t); in val cv = @{computation_conv "pexpr \ pexpr \ pexpr list" terms: fnorm nat_of_integer Code_Target_Nat.natural "0::nat" "1::nat" "2::nat" "3::nat" "0::int" "1::int" "2::int" "3::int" "-1::int" datatypes: fexpr int integer num} (fn ctxt => fn result => fn ct => fnorm_oracle ctxt ct (term_of_result result)) end \ ML \ signature FIELD_TAC = sig structure Field_Simps: sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end val eq_field_simps: (term * (thm list * thm list * thm list * thm * thm)) * (term * (thm list * thm list * thm list * thm * thm)) -> bool val field_tac: bool -> Proof.context -> int -> tactic end structure Field_Tac : FIELD_TAC = struct open Ring_Tac; fun field_struct \<^Const_>\Ring.ring.add _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Ring.a_minus _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Group.monoid.mult _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Ring.a_inv _ _ for R _\ = SOME R | field_struct \<^Const_>\Group.pow _ _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Algebra_Aux.m_div _ _for R _ _\ = SOME R | field_struct \<^Const_>\Ring.ring.zero _ _ for R\ = SOME R | field_struct \<^Const_>\Group.monoid.one _ _ for R\ = SOME R | field_struct \<^Const_>\Algebra_Aux.of_integer _ _ for R _\ = SOME R | field_struct _ = NONE; fun reif_fexpr vs \<^Const_>\Ring.ring.add _ _ for _ a b\ = \<^Const>\FAdd for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Ring.a_minus _ _ for _ a b\ = \<^Const>\FSub for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Group.monoid.mult _ _ for _ a b\ = \<^Const>\FMul for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Ring.a_inv _ _ for _ a\ = \<^Const>\FNeg for \reif_fexpr vs a\\ | reif_fexpr vs \<^Const>\Group.pow _ _ _ for _ a n\ = \<^Const>\FPow for \reif_fexpr vs a\ n\ | reif_fexpr vs \<^Const_>\Algebra_Aux.m_div _ _ for _ a b\ = \<^Const>\FDiv for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs (Free x) = \<^Const>\FVar for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_fexpr vs \<^Const_>\Ring.ring.zero _ _ for _\ = \<^term>\FCnst 0\ | reif_fexpr vs \<^Const_>\Group.monoid.one _ _ for _\ = \<^term>\FCnst 1\ | reif_fexpr vs \<^Const_>\Algebra_Aux.of_integer _ _ for _ n\ = \<^Const>\FCnst for n\ | reif_fexpr _ _ = error "reif_fexpr: bad expression"; fun reif_fexpr' vs \<^Const_>\plus _ for a b\ = \<^Const>\FAdd for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\minus _ for a b\ = \<^Const>\FSub for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\times _ for a b\ = \<^Const>\FMul for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\uminus _ for a\ = \<^Const>\FNeg for \reif_fexpr' vs a\\ | reif_fexpr' vs \<^Const_>\power _ for a n\ = \<^Const>\FPow for \reif_fexpr' vs a\ n\ | reif_fexpr' vs \<^Const_>\divide _ for a b\ = \<^Const>\FDiv for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs (Free x) = \<^Const>\FVar for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_fexpr' vs \<^Const_>\zero_class.zero _\ = \<^term>\FCnst 0\ | reif_fexpr' vs \<^Const_>\one_class.one _\ = \<^term>\FCnst 1\ | reif_fexpr' vs \<^Const_>\numeral _ for b\ = \<^Const>\FCnst for \<^Const>\numeral \<^Type>\int\ for b\\ | reif_fexpr' _ _ = error "reif_fexpr: bad expression"; fun eq_field_simps ((t, (ths1, ths2, ths3, th4, th)), (t', (ths1', ths2', ths3', th4', th'))) = t aconv t' andalso eq_list Thm.eq_thm (ths1, ths1') andalso eq_list Thm.eq_thm (ths2, ths2') andalso eq_list Thm.eq_thm (ths3, ths3') andalso Thm.eq_thm (th4, th4') andalso Thm.eq_thm (th, th'); structure Field_Simps = Generic_Data (struct type T = (term * (thm list * thm list * thm list * thm * thm)) Net.net val empty = Net.empty val merge = Net.merge eq_field_simps end); fun get_field_simps ctxt optcT t = (case get_matching_rules ctxt (Field_Simps.get (Context.Proof ctxt)) t of SOME (ths1, ths2, ths3, th4, th) => let val tr = Thm.transfer' ctxt #> (case optcT of NONE => I | SOME cT => inst [cT] [] #> norm) in (map tr ths1, map tr ths2, map tr ths3, tr th4, tr th) end | NONE => error "get_field_simps: lookup failed"); fun nth_el_conv (_, _, _, nth_el_Cons, _) = let val a = type_of_eqn nth_el_Cons; val If_conv_a = If_conv a; fun conv ys n = (case strip_app ys of (\<^const_name>\Cons\, [x, xs]) => transitive' (inst [] [x, xs, n] nth_el_Cons) (If_conv_a (args2 nat_eq_conv) Thm.reflexive (cong2' conv Thm.reflexive (args2 nat_minus_conv)))) in conv end; fun feval_conv (rls as ([feval_simps_1, feval_simps_2, feval_simps_3, feval_simps_4, feval_simps_5, feval_simps_6, feval_simps_7, feval_simps_8, feval_simps_9, feval_simps_10, feval_simps_11], _, _, _, _)) = let val nth_el_conv' = nth_el_conv rls; fun conv xs x = (case strip_app x of (\<^const_name>\FCnst\, [c]) => (case strip_app c of (\<^const_name>\zero_class.zero\, _) => inst [] [xs] feval_simps_9 | (\<^const_name>\one_class.one\, _) => inst [] [xs] feval_simps_10 | (\<^const_name>\numeral\, [n]) => inst [] [xs, n] feval_simps_11 | _ => inst [] [xs, c] feval_simps_1) | (\<^const_name>\FVar\, [n]) => transitive' (inst [] [xs, n] feval_simps_2) (args2 nth_el_conv') | (\<^const_name>\FAdd\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_3) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FSub\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_4) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FMul\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_5) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FNeg\, [a]) => transitive' (inst [] [xs, a] feval_simps_6) (cong1 (args2 conv)) | (\<^const_name>\FDiv\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_7) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FPow\, [a, n]) => transitive' (inst [] [xs, a, n] feval_simps_8) (cong2 (args2 conv) Thm.reflexive)) in conv end; fun peval_conv (rls as (_, [peval_simps_1, peval_simps_2, peval_simps_3, peval_simps_4, peval_simps_5, peval_simps_6, peval_simps_7, peval_simps_8, peval_simps_9, peval_simps_10, peval_simps_11], _, _, _)) = let val nth_el_conv' = nth_el_conv rls; fun conv xs x = (case strip_app x of (\<^const_name>\PExpr1\, [e]) => (case strip_app e of (\<^const_name>\PCnst\, [c]) => (case strip_numeral c of (\<^const_name>\zero_class.zero\, _) => inst [] [xs] peval_simps_8 | (\<^const_name>\one_class.one\, _) => inst [] [xs] peval_simps_9 | (\<^const_name>\numeral\, [n]) => inst [] [xs, n] peval_simps_10 | (\<^const_name>\uminus\, [n]) => inst [] [xs, n] peval_simps_11 | _ => inst [] [xs, c] peval_simps_1) | (\<^const_name>\PVar\, [n]) => transitive' (inst [] [xs, n] peval_simps_2) (args2 nth_el_conv') | (\<^const_name>\PAdd\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_3) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PSub\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_4) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PNeg\, [a]) => transitive' (inst [] [xs, a] peval_simps_5) (cong1 (args2 conv))) | (\<^const_name>\PExpr2\, [e]) => (case strip_app e of (\<^const_name>\PMul\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_6) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PPow\, [a, n]) => transitive' (inst [] [xs, a, n] peval_simps_7) (cong2 (args2 conv) Thm.reflexive))) in conv end; fun nonzero_conv (rls as (_, _, [nonzero_Nil, nonzero_Cons, nonzero_singleton], _, _)) = let val peval_conv' = peval_conv rls; fun conv xs qs = (case strip_app qs of (\<^const_name>\Nil\, []) => inst [] [xs] nonzero_Nil | (\<^const_name>\Cons\, [p, ps]) => (case Thm.term_of ps of \<^Const_>\Nil _\ => transitive' (inst [] [xs, p] nonzero_singleton) (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) | _ => transitive' (inst [] [xs, p, ps] nonzero_Cons) (cong2 (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) (args2 conv)))) in conv end; fun field_tac in_prem ctxt = SUBGOAL (fn (g, i) => let val (prems, concl) = Logic.strip_horn g; fun find_eq s = (case s of (_ $ \<^Const_>\HOL.eq T for t u\) => (case (field_struct t, field_struct u) of (SOME R, _) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr) | (_, SOME R) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr) | _ => if Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\field\) then SOME ((t, u), mk_ring T, T, SOME T, K @{thm in_carrier_trivial}, reif_fexpr') else NONE) | _ => NONE); val ((t, u), R, T, optT, mkic, reif) = (case get_first find_eq (if in_prem then prems else [concl]) of SOME q => q | NONE => error "cannot determine field"); val rls as (_, _, _, _, feval_eq) = get_field_simps ctxt (Option.map (Thm.ctyp_of ctxt) optT) R; val xs = [] |> Term.add_frees t |> Term.add_frees u |> filter (equal T o snd); val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs)); val ce = Thm.cterm_of ctxt (reif xs t); val ce' = Thm.cterm_of ctxt (reif xs u); val fnorm = cv ctxt \<^instantiate>\e = ce and e' = ce' in cterm \fnorm (FSub e e')\\; val (_, [n, dc]) = strip_app (Thm.rhs_of fnorm); val (_, [_, c]) = strip_app dc; val th = Conv.fconv_rule (Conv.concl_conv 1 (Conv.arg_conv (binop_conv (binop_conv (K (feval_conv rls cxs ce)) (K (feval_conv rls cxs ce'))) (Conv.arg1_conv (K (peval_conv rls cxs n)))))) ([mkic xs, HOLogic.mk_obj_eq fnorm, HOLogic.mk_obj_eq (nonzero_conv rls cxs c) RS @{thm iffD2}] MRS feval_eq); val th' = Drule.rotate_prems 1 (th RS (if in_prem then @{thm iffD1} else @{thm iffD2})); in if in_prem then dresolve_tac ctxt [th'] 1 THEN defer_tac 1 else resolve_tac ctxt [th'] 1 end); end \ context field begin local_setup \ -Local_Theory.declaration {syntax = false, pervasive = false} +Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps (Morphism.term phi \<^term>\R\, (Morphism.fact phi @{thms feval.simps [meta] feval_Cnst [meta]}, Morphism.fact phi @{thms peval.simps [meta] peval_Cnst [meta]}, Morphism.fact phi @{thms nonzero.simps [meta] nonzero_singleton [meta]}, singleton (Morphism.fact phi) @{thm nth_el_Cons [meta]}, singleton (Morphism.fact phi) @{thm feval_eq})))) \ end method_setup field = \ Scan.lift (Args.mode "prems") -- Attrib.thms >> (fn (in_prem, thms) => fn ctxt => SIMPLE_METHOD' (Field_Tac.field_tac in_prem ctxt THEN' Ring_Tac.ring_tac in_prem thms ctxt)) \ "reduce equations over fields to equations over rings" end diff --git a/src/HOL/Eisbach/Tests.thy b/src/HOL/Eisbach/Tests.thy --- a/src/HOL/Eisbach/Tests.thy +++ b/src/HOL/Eisbach/Tests.thy @@ -1,606 +1,606 @@ (* Title: HOL/Eisbach/Tests.thy Author: Daniel Matichuk, NICTA/UNSW *) section \Miscellaneous Eisbach tests\ theory Tests imports Main Eisbach_Tools begin subsection \Named Theorems Tests\ named_theorems foo method foo declares foo = (rule foo) lemma assumes A [foo]: A shows A apply foo done method abs_used for P = (match (P) in "\a. ?Q" \ fail \ _ \ -) subsection \Match Tests\ notepad begin have dup: "\A. A \ A \ A" by simp fix A y have "(\x. A x) \ A y" apply (rule dup, match premises in Y: "\B. P B" for P \ \match (P) in A \ \print_fact Y, rule Y\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match (P) in A \ \print_fact Y, rule Y\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match conclusion in "P y" for y \ \print_fact Y, print_term y, rule Y[where B=y]\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match conclusion in "P z" for z \ \print_fact Y, print_term y, rule Y[where B=z]\\) apply (rule dup, match conclusion in "P y" for P \ \match premises in Y: "\z. P z" \ \print_fact Y, rule Y[where z=y]\\) apply (match premises in Y: "\z :: 'a. P z" for P \ \match conclusion in "P y" \ \print_fact Y, rule Y[where z=y]\\) done assume X: "\x. A x" "A y" have "A y" apply (match X in Y:"\B. A B" and Y':"B ?x" for B \ \print_fact Y[where B=y], print_term B\) apply (match X in Y:"B ?x" and Y':"B ?x" for B \ \print_fact Y, print_term B\) apply (match X in Y:"B x" and Y':"B x" for B x \ \print_fact Y, print_term B, print_term x\) apply (insert X) apply (match premises in Y:"\B. A B" and Y':"B y" for B and y :: 'a \ \print_fact Y[where B=y], print_term B\) apply (match premises in Y:"B ?x" and Y':"B ?x" for B \ \print_fact Y, print_term B\) apply (match premises in Y:"B x" and Y':"B x" for B x \ \print_fact Y, print_term B\) apply (match conclusion in "P x" and "P y" for P x \ \print_term P, print_term x\) apply assumption done { fix B x y assume X: "\x y. B x x y" have "B x x y" by (match X in Y:"\y. B y y z" for z \ \rule Y[where y=x]\) fix A B have "(\x y. A (B x) y) \ A (B x) y" by (match premises in Y: "\xx. ?H (B xx)" \ \rule Y\) } (* match focusing retains prems *) fix B x have "(\x. A x) \ (\z. B z) \ A y \ B x" apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y': "\z :: 'b. B z" \ \print_fact Y, print_fact Y', rule Y'[where z=x]\\) done (*Attributes *) fix C have "(\x :: 'a. A x) \ (\z. B z) \ A y \ B x \ B x \ A y" apply (intro conjI) apply (match premises in Y: "\z :: 'a. A z" and Y'[intro]:"\z :: 'b. B z" \ fastforce) apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y'[intro]:"\z :: 'b. B z" \ fastforce\) apply (match premises in Y[thin]: "\z :: 'a. A z" \ \(match premises in Y':"\z :: 'a. A z" \ \print_fact Y,fail\ \ _ \ \print_fact Y\)\) (*apply (match premises in Y: "\z :: 'b. B z" \ \(match premises in Y'[thin]:"\z :: 'b. B z" \ \(match premises in Y':"\z :: 'a. A z" \ fail \ Y': _ \ -)\)\)*) apply assumption done fix A B C D have "\uu'' uu''' uu uu'. (\x :: 'a. A uu' x) \ D uu y \ (\z. B uu z) \ C uu y \ (\z y. C uu z) \ B uu x \ B uu x \ C uu y" apply (match premises in Y[thin]: "\z :: 'a. A ?zz' z" and Y'[thin]: "\rr :: 'b. B ?zz rr" \ \print_fact Y, print_fact Y', intro conjI, rule Y', insert Y', insert Y'[where rr=x]\) apply (match premises in Y:"B ?u ?x" \ \rule Y\) apply (insert TrueI) apply (match premises in Y'[thin]: "\ff. B uu ff" for uu \ \insert Y', drule meta_spec[where x=x]\) apply assumption done (* Multi-matches. As many facts as match are bound. *) fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ \intro conjI, (rule Y)+\) apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z. ?A z" (multi) \ \intro conjI, (rule Y)+\) apply (match premises in Y[thin]: "\z. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done fix A B C P Q and x :: 'a and y :: 'a have "(\x y :: 'a. A x y \ Q) \ (\a b. B (a :: 'a) (b :: 'a) \ Q) \ (\x y. C (x :: 'a) (y :: 'a) \ P) \ A y x \ B y x" by (match premises in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule conjI, rule Y[where z=x,THEN conjunct1], rule Y[THEN conjunct1]\) (*We may use for-fixes in multi-matches too. All bound facts must agree on the fixed term *) fix A B C x have "(\y :: 'a. B y \ C y) \ (\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match (P) in B \ fail \ "\a. B" \ fail \ _ \ -, intro conjI, (rule Y[THEN conjunct1])\) apply (rule dup) apply (match premises in Y':"\x :: 'a. ?U x \ Q x" and Y: "\x :: 'a. Q x \ ?U x" (multi) for Q \ \insert Y[THEN conjunct1]\) apply assumption (* Previous match requires that Q is consistent *) apply (match premises in Y: "\z :: 'a. ?A z \ False" (multi) \ \print_fact Y, fail\ \ "C y" \ \print_term C\) (* multi-match must bind something *) apply (match premises in Y: "\x. B x \ C x" \ \intro conjI Y[THEN conjunct1]\) apply (match premises in Y: "C ?x" \ \rule Y\) done (* All bindings must be tried for a particular theorem. However all combinations are NOT explored. *) fix B A C assume asms:"\a b. B (a :: 'a) (b :: 'a) \ Q" "\x :: 'a. A x x \ Q" "\a b. C (a :: 'a) (b :: 'a) \ Q" have "B y x \ C x y \ B x y \ C y x \ A x x" apply (intro conjI) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where z=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where z=x,THEN conjunct1]\) apply (match asms in Y: "\z a. A (z :: 'a) (a :: 'a) \ R" for R \ fail \ _ \ -) apply (rule asms[THEN conjunct1]) done (* Attributes *) fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match Y[THEN conjunct1] in Y':"?H" (multi) \ \intro conjI,rule Y'\\) apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match Y[THEN conjunct2] in Y':"?H" (multi) \ \rule Y'\\) apply assumption done (* Removed feature for now *) (* fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match prems in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match \K @{thms Y TrueI}\ in Y':"?H" (multi) \ \rule conjI; (rule Y')?\\) apply (match prems in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match \K [@{thm Y}]\ in Y':"?H" (multi) \ \rule Y'\\) done *) (* Testing THEN_ALL_NEW within match *) fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \intro conjI ; ((rule Y[THEN conjunct1])?); rule Y[THEN conjunct2] \) done (* Cut tests *) fix A B C D have "D \ C \ A \ B \ A \ C \ D \ True \ C" by (((match premises in I: "P \ Q" (cut) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\)?), simp) have "D \ C \ A \ B \ A \ C \ D \ True \ C" by (match premises in I: "P \ Q" (cut 2) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\) have "A \ B \ A \ C \ C" by (((match premises in I: "P \ Q" (cut) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\)?, simp) | simp) fix f x y have "f x y \ f x y" by (match conclusion in "f x y" for f x y \ \print_term f\) fix A B C assume X: "A \ B" "A \ C" C have "A \ B \ C" by (match X in H: "A \ ?H" (multi, cut) \ \match H in "A \ C" and "A \ B" \ fail\ | simp add: X) (* Thinning an inner focus *) (* Thinning should persist within a match, even when on an external premise *) fix A have "(\x. A x \ B) \ B \ C \ C" apply (match premises in H:"\x. A x \ B" \ \match premises in H'[thin]: "\x. A x \ B" \ \match premises in H'':"\x. A x \ B" \ fail \ _ \ -\ ,match premises in H'':"\x. A x \ B" \ fail \ _ \ -\) apply (match premises in H:"\x. A x \ B" \ fail \ H':_ \ \rule H'[THEN conjunct2]\) done (* Local premises *) (* Only match premises which actually existed in the goal we just focused.*) fix A assume asms: "C \ D" have "B \ C \ C" by (match premises in _ \ \insert asms, match premises (local) in "B \ C" \ fail \ H:"C \ D" \ \rule H[THEN conjunct1]\\) end (* Testing inner focusing. This fails if we don't smash flex-flex pairs produced by retrofitting. This needs to be done more carefully to avoid smashing legitimate pairs.*) schematic_goal "?A x \ A x" apply (match conclusion in "H" for H \ \match conclusion in Y for Y \ \print_term Y\\) apply assumption done (* Ensure short-circuit after first match failure *) lemma assumes A: "P \ Q" shows "P" by ((match A in "P \ Q" \ fail \ "?H" \ -) | simp add: A) lemma assumes A: "D \ C" "A \ B" "A \ B" shows "A" apply ((match A in U: "P \ Q" (cut) and "P' \ Q'" for P Q P' Q' \ \simp add: U\ \ "?H" \ -) | -) apply (simp add: A) done subsection \Uses Tests\ ML \ fun test_internal_fact ctxt factnm = (case \<^try>\Proof_Context.get_thms ctxt factnm\ of NONE => () | SOME _ => error "Found internal fact"); \ method uses_test\<^sub>1 uses uses_test\<^sub>1_uses = (rule uses_test\<^sub>1_uses) lemma assumes A shows A by (uses_test\<^sub>1 uses_test\<^sub>1_uses: assms) ML \test_internal_fact \<^context> "uses_test\<^sub>1_uses"\ ML \test_internal_fact \<^context> "Tests.uses_test\<^sub>1_uses"\ ML \test_internal_fact \<^context> "Tests.uses_test\<^sub>1.uses_test\<^sub>1_uses"\ subsection \Basic fact passing\ method find_fact for x y :: bool uses facts1 facts2 = (match facts1 in U: "x" \ \insert U, match facts2 in U: "y" \ \insert U\\) lemma assumes A: A and B: B shows "A \ B" apply (find_fact "A" "B" facts1: A facts2: B) apply (rule conjI; assumption) done subsection \Testing term and fact passing in recursion\ method recursion_example for x :: bool uses facts = (match (x) in "A \ B" for A B \ \(recursion_example A facts: facts, recursion_example B facts: facts)\ \ "?H" \ \match facts in U: "x" \ \insert U\\) lemma assumes asms: "A" "B" "C" "D" shows "(A \ B) \ (C \ D)" apply (recursion_example "(A \ B) \ (C \ D)" facts: asms) apply simp done (* uses facts are not accumulated *) method recursion_example' for A :: bool and B :: bool uses facts = (match facts in H: "A" and H': "B" \ \recursion_example' "A" "B" facts: H TrueI\ \ "A" and "True" \ \recursion_example' "A" "B" facts: TrueI\ \ "True" \ - \ "PROP ?P" \ fail) lemma assumes asms: "A" "B" shows "True" apply (recursion_example' "A" "B" facts: asms) apply simp done (*Method.sections in existing method*) method my_simp\<^sub>1 uses my_simp\<^sub>1_facts = (simp add: my_simp\<^sub>1_facts) lemma assumes A shows A by (my_simp\<^sub>1 my_simp\<^sub>1_facts: assms) (*Method.sections via Eisbach argument parser*) method uses_test\<^sub>2 uses uses_test\<^sub>2_uses = (uses_test\<^sub>1 uses_test\<^sub>1_uses: uses_test\<^sub>2_uses) lemma assumes A shows A by (uses_test\<^sub>2 uses_test\<^sub>2_uses: assms) subsection \Declaration Tests\ named_theorems declare_facts\<^sub>1 method declares_test\<^sub>1 declares declare_facts\<^sub>1 = (rule declare_facts\<^sub>1) lemma assumes A shows A by (declares_test\<^sub>1 declare_facts\<^sub>1: assms) lemma assumes A[declare_facts\<^sub>1]: A shows A by declares_test\<^sub>1 subsection \Rule Instantiation Tests\ method my_allE\<^sub>1 for x :: 'a and P :: "'a \ bool" = (erule allE [where x = x and P = P]) lemma "\x. Q x \ Q x" by (my_allE\<^sub>1 x Q) method my_allE\<^sub>2 for x :: 'a and P :: "'a \ bool" = (erule allE [of P x]) lemma "\x. Q x \ Q x" by (my_allE\<^sub>2 x Q) method my_allE\<^sub>3 for x :: 'a and P :: "'a \ bool" = (match allE [where 'a = 'a] in X: "\(x :: 'a) P R. \x. P x \ (P x \ R) \ R" \ \erule X [where x = x and P = P]\) lemma "\x. Q x \ Q x" by (my_allE\<^sub>3 x Q) method my_allE\<^sub>4 for x :: 'a and P :: "'a \ bool" = (match allE [where 'a = 'a] in X: "\(x :: 'a) P R. \x. P x \ (P x \ R) \ R" \ \erule X [of x P]\) lemma "\x. Q x \ Q x" by (my_allE\<^sub>4 x Q) subsection \Polymorphism test\ axiomatization foo' :: "'a \ 'b \ 'c \ bool" axiomatization where foo'_ax1: "foo' x y z \ z \ y" axiomatization where foo'_ax2: "foo' x y y \ x \ z" axiomatization where foo'_ax3: "foo' (x :: int) y y \ y \ y" lemmas my_thms = foo'_ax1 foo'_ax2 foo'_ax3 definition first_id where "first_id x = x" lemmas my_thms' = my_thms[of "first_id x" for x] method print_conclusion = (match conclusion in concl for concl \ \print_term concl\) lemma assumes foo: "\x (y :: bool). foo' (A x) B (A x)" shows "\z. A z \ B" apply (match conclusion in "f x y" for f y and x :: "'d :: type" \ \ match my_thms' in R:"\(x :: 'f :: type). ?P (first_id x) \ ?R" and R':"\(x :: 'f :: type). ?P' (first_id x) \ ?R'" \ \ match (x) in "q :: 'f" for q \ \ rule R[of q,simplified first_id_def], print_conclusion, rule foo \\\) done subsection \Unchecked rule instantiation, with the possibility of runtime errors\ named_theorems my_thms_named declare foo'_ax3[my_thms_named] method foo_method3 declares my_thms_named = (match my_thms_named[of (unchecked) z for z] in R:"PROP ?H" \ \rule R\) notepad begin (*FIXME: Shouldn't need unchecked keyword here. See Tests_Failing.thy *) fix A B x have "foo' x B A \ A \ B" by (match my_thms[of (unchecked) z for z] in R:"PROP ?H" \ \rule R\) fix A B x note foo'_ax1[my_thms_named] have "foo' x B A \ A \ B" by (match my_thms_named[where x=z for z] in R:"PROP ?H" \ \rule R\) fix A B x note foo'_ax1[my_thms_named] foo'_ax2[my_thms_named] foo'_ax3[my_thms_named] have "foo' x B A \ A \ B" by foo_method3 end ML \ structure Data = Generic_Data ( type T = thm list; val empty: T = []; fun merge data : T = Thm.merge_thms data; ); \ local_setup \Local_Theory.add_thms_dynamic (\<^binding>\test_dyn\, Data.get)\ setup \Context.theory_map (Data.put @{thms TrueI})\ method dynamic_thms_test = (rule test_dyn) locale foo = fixes A assumes A : "A" begin local_setup - \Local_Theory.declaration {pervasive = false, syntax = false} + \Local_Theory.declaration {pervasive = false, syntax = false, pos = \<^here>} (fn phi => Data.put (Morphism.fact phi @{thms A}))\ lemma A by dynamic_thms_test end notepad begin fix A x assume X: "\x. A x" have "A x" by (match X in H[of x]:"\x. A x" \ \print_fact H,match H in "A x" \ \rule H\\) fix A x B assume X: "\x :: bool. A x \ B" "\x. A x" assume Y: "A B" have "B \ B \ B \ B \ B \ B" apply (intro conjI) apply (match X in H[OF X(2)]:"\x. A x \ B" \ \print_fact H,rule H\) apply (match X in H':"\x. A x" and H[OF H']:"\x. A x \ B" \ \print_fact H',print_fact H,rule H\) apply (match X in H[of Q]:"\x. A x \ ?R" and "?P \ Q" for Q \ \print_fact H,rule H, rule Y\) apply (match X in H[of Q,OF Y]:"\x. A x \ ?R" and "?P \ Q" for Q \ \print_fact H,rule H\) apply (match X in H[OF Y,intro]:"\x. A x \ ?R" \ \print_fact H,fastforce\) apply (match X in H[intro]:"\x. A x \ ?R" \ \rule H[where x=B], rule Y\) done fix x :: "prop" and A assume X: "TERM x" assume Y: "\x :: prop. A x" have "A TERM x" apply (match X in "PROP y" for y \ \rule Y[where x="PROP y"]\) done end subsection \Proper context for method parameters\ method add_simp methods m uses f = (match f in H[simp]:_ \ m) method add_my_thms methods m uses f = (match f in H[my_thms_named]:_ \ m) method rule_my_thms = (rule my_thms_named) method rule_my_thms' declares my_thms_named = (rule my_thms_named) lemma assumes A: A and B: B shows "(A \ B) \ A \ A \ A" apply (intro conjI) apply (add_simp \add_simp simp f: B\ f: A) apply (add_my_thms rule_my_thms f:A) apply (add_my_thms rule_my_thms' f:A) apply (add_my_thms \rule my_thms_named\ f:A) done subsection \Shallow parser tests\ method all_args for A B methods m1 m2 uses f1 f2 declares my_thms_named = fail lemma True by (all_args True False - fail f1: TrueI f2: TrueI my_thms_named: TrueI | rule TrueI) subsection \Method name internalization test\ method test2 = (simp) method simp = fail lemma "A \ A" by test2 subsection \Dynamic facts\ named_theorems my_thms_named' method foo_method1 for x = (match my_thms_named' [of (unchecked) x] in R: "PROP ?H" \ \rule R\) lemma assumes A [my_thms_named']: "\x. A x" shows "A y" by (foo_method1 y) subsection \Eisbach method invocation from ML\ method test_method for x y uses r = (print_term x, print_term y, rule r) method_setup test_method' = \ Args.term -- Args.term -- (Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms) >> (fn ((x, y), r) => fn ctxt => Method_Closure.apply_method ctxt \<^method>\test_method\ [x, y] [r] [] ctxt) \ lemma fixes a b :: nat assumes "a = b" shows "b = a" apply (test_method a b)? apply (test_method' a b rule: refl)? apply (test_method' a b rule: assms [symmetric])? done subsection \Eisbach methods in locales\ locale my_locale1 = fixes A assumes A: A begin method apply_A = (match conclusion in "A" \ \match A in U:"A" \ \print_term A, print_fact A, rule U\\) end locale my_locale2 = fixes B assumes B: B begin interpretation my_locale1 B by (unfold_locales; rule B) lemma B by apply_A end context fixes C assumes C: C begin interpretation my_locale1 C by (unfold_locales; rule C) lemma C by apply_A end context begin interpretation my_locale1 "True \ True" by (unfold_locales; blast) lemma "True \ True" by apply_A end locale locale_poly = fixes P assumes P: "\x :: 'a. P x" begin method solve_P for z :: 'a = (rule P[where x = z]) end context begin interpretation locale_poly "\x:: nat. 0 \ x" by (unfold_locales; blast) lemma "0 \ (n :: nat)" by (solve_P n) end subsection \Mutual recursion via higher-order methods\ experiment begin method inner_method methods passed_method = (rule conjI; passed_method) method outer_method = (inner_method \outer_method\ | assumption) lemma "Q \ R \ P \ (Q \ R) \ P" by outer_method end end diff --git a/src/HOL/Eisbach/method_closure.ML b/src/HOL/Eisbach/method_closure.ML --- a/src/HOL/Eisbach/method_closure.ML +++ b/src/HOL/Eisbach/method_closure.ML @@ -1,251 +1,252 @@ (* Title: HOL/Eisbach/method_closure.ML Author: Daniel Matichuk, NICTA/UNSW Facilities for treating method syntax as a closure, with abstraction over terms, facts and other methods. The 'method' command allows to define new proof methods by combining existing ones with their usual syntax. *) signature METHOD_CLOSURE = sig val apply_method: Proof.context -> string -> term list -> thm list list -> (Proof.context -> Method.method) list -> Proof.context -> thm list -> context_tactic val method: binding -> (binding * typ option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory val method_cmd: binding -> (binding * string option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory end; structure Method_Closure: METHOD_CLOSURE = struct (* auxiliary data for method definition *) structure Method_Definition = Proof_Data ( type T = (Proof.context -> Method.method) Symtab.table * (*dynamic methods*) (term list -> Proof.context -> Method.method) Symtab.table (*recursive methods*); fun init _ : T = (Symtab.empty, Symtab.empty); ); fun lookup_dynamic_method ctxt full_name = (case Symtab.lookup (#1 (Method_Definition.get ctxt)) full_name of SOME m => m ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val update_dynamic_method = Method_Definition.map o apfst o Symtab.update; fun get_recursive_method full_name ts ctxt = (case Symtab.lookup (#2 (Method_Definition.get ctxt)) full_name of SOME m => m ts ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val put_recursive_method = Method_Definition.map o apsnd o Symtab.update; (* stored method closures *) type closure = {vars: term list, named_thms: string list, methods: string list, body: Method.text}; structure Data = Generic_Data ( type T = closure Symtab.table; val empty: T = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun get_closure ctxt name = (case Symtab.lookup (Data.get (Context.Proof ctxt)) name of SOME closure => closure | NONE => error ("Unknown Eisbach method: " ^ quote name)); fun put_closure binding (closure: closure) lthy = let val name = Local_Theory.full_name lthy binding; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map - (Symtab.update (name, - {vars = map (Morphism.term phi) (#vars closure), - named_thms = #named_thms closure, - methods = #methods closure, - body = (Method.map_source o map) (Token.transform phi) (#body closure)}))) + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding} + (fn phi => + Data.map + (Symtab.update (name, + {vars = map (Morphism.term phi) (#vars closure), + named_thms = #named_thms closure, + methods = #methods closure, + body = (Method.map_source o map) (Token.transform phi) (#body closure)}))) end; (* instantiate and evaluate method text *) fun method_instantiate vars body ts ctxt = let val thy = Proof_Context.theory_of ctxt; val subst = fold (Pattern.match thy) (vars ~~ ts) (Vartab.empty, Vartab.empty); val morphism = Morphism.term_morphism "method_instantiate" (Envir.subst_term subst); val body' = (Method.map_source o map) (Token.transform morphism) body; in Method.evaluate_runtime body' ctxt end; (** apply method closure **) fun recursive_method full_name vars body ts = let val m = method_instantiate vars body in put_recursive_method (full_name, m) #> m ts end; fun apply_method ctxt method_name terms facts methods = let fun declare_facts (name :: names) (fact :: facts) = fold (Context.proof_map o Named_Theorems.add_thm name) fact #> declare_facts names facts | declare_facts _ [] = I | declare_facts [] (_ :: _) = error ("Excessive facts for method " ^ quote method_name); val {vars, named_thms, methods = method_args, body} = get_closure ctxt method_name; in declare_facts named_thms facts #> fold update_dynamic_method (method_args ~~ methods) #> recursive_method method_name vars body terms end; (** define method closure **) local fun setup_local_method binding lthy = let val full_name = Local_Theory.full_name lthy binding; fun dynamic_method ctxt = lookup_dynamic_method ctxt full_name; in lthy |> update_dynamic_method (full_name, K Method.fail) |> Method.local_setup binding (Scan.succeed dynamic_method) "(internal)" end; fun check_named_thm ctxt binding = let val bname = Binding.name_of binding; val pos = Binding.pos_of binding; val full_name = Named_Theorems.check ctxt (bname, pos); val parser: Method.modifier parser = Args.$$$ bname -- Args.colon >> K {init = I, attribute = Named_Theorems.add full_name, pos = pos}; in (full_name, parser) end; fun parse_term_args args = Args.context :|-- (fn ctxt => let val ctxt' = Proof_Context.set_mode (Proof_Context.mode_schematic) ctxt; fun parse T = (if T = propT then Syntax.parse_prop ctxt' else Syntax.parse_term ctxt') #> Type.constraint (Type_Infer.paramify_vars T); fun do_parse' T = Parse_Tools.name_term >> Parse_Tools.parse_val_cases (parse T); fun do_parse (Var (_, T)) = do_parse' T | do_parse (Free (_, T)) = do_parse' T | do_parse t = error ("Unexpected method parameter: " ^ Syntax.string_of_term ctxt' t); fun rep [] x = Scan.succeed [] x | rep (t :: ts) x = (do_parse t ::: rep ts) x; fun check ts = let val (ts, fs) = split_list ts; val ts' = Syntax.check_terms ctxt' ts |> Variable.polymorphic ctxt'; val _ = ListPair.app (fn (f, t) => f t) (fs, ts'); in ts' end; in Scan.lift (rep args) >> check end); fun parse_method_args method_args = let fun bind_method (name, text) ctxt = let val method = Method.evaluate_runtime text; val inner_update = method o update_dynamic_method (name, K (method ctxt)); in update_dynamic_method (name, inner_update) ctxt end; fun rep [] x = Scan.succeed [] x | rep (m :: ms) x = ((Method.text_closure >> pair m) ::: rep ms) x; in rep method_args >> fold bind_method end; fun gen_method add_fixes name vars uses declares methods source lthy = let val (uses_internal, lthy1) = lthy |> Proof_Context.concealed |> Local_Theory.begin_nested |-> Proof_Context.private_scope |> Local_Theory.map_background_naming (Name_Space.add_path (Binding.name_of name)) |> Config.put Method.old_section_parser true |> fold setup_local_method methods |> fold_map (fn b => Named_Theorems.declare b "") uses; val (term_args, lthy2) = lthy1 |> add_fixes vars |-> fold_map Proof_Context.inferred_param |>> map Free; val (named_thms, modifiers) = map (check_named_thm lthy2) (declares @ uses) |> split_list; val method_args = map (Local_Theory.full_name lthy2) methods; fun parser args meth = apfst (Config.put_generic Method.old_section_parser true) #> (parse_term_args args -- parse_method_args method_args --| (Scan.depend (fn context => Scan.succeed (fold Named_Theorems.clear uses_internal context, ())) -- Method.sections modifiers)) >> (fn (ts, decl) => meth ts o decl); val full_name = Local_Theory.full_name lthy name; val lthy3 = lthy2 |> Method.local_setup (Binding.make (Binding.name_of name, Position.none)) (parser term_args (get_recursive_method full_name)) "(internal)" |> put_recursive_method (full_name, fn _ => fn _ => Method.fail); val (text, src) = Method.read_closure (Config.put Proof_Context.dynamic_facts_dummy true lthy3) source; val morphism = Variable.export_morphism lthy3 (lthy |> Proof_Context.transfer (Proof_Context.theory_of lthy3) |> fold Token.declare_maxidx src |> Variable.declare_maxidx (Variable.maxidx_of lthy3)); val text' = (Method.map_source o map) (Token.transform morphism) text; val term_args' = map (Morphism.term morphism) term_args; in lthy3 |> Local_Theory.end_nested |> Proof_Context.restore_naming lthy |> put_closure name {vars = term_args', named_thms = named_thms, methods = method_args, body = text'} |> Method.local_setup name (Args.context :|-- (fn ctxt => let val {body, vars, ...} = get_closure ctxt full_name in parser vars (recursive_method full_name vars body) end)) "" |> pair full_name end; in val method = gen_method Proof_Context.add_fixes; val method_cmd = gen_method Proof_Context.add_fixes_cmd; end; val _ = Outer_Syntax.local_theory \<^command_keyword>\method\ "Eisbach method definition" (Parse.binding -- Parse.for_fixes -- ((Scan.optional (\<^keyword>\methods\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- (Scan.optional (\<^keyword>\uses\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) [])) -- (Scan.optional (\<^keyword>\declares\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- Parse.!!! (\<^keyword>\=\ |-- Parse.args1 (K true)) >> (fn ((((name, vars), (methods, uses)), declares), src) => #2 o method_cmd name vars uses declares methods src)); end; diff --git a/src/HOL/HOLCF/Tools/fixrec.ML b/src/HOL/HOLCF/Tools/fixrec.ML --- a/src/HOL/HOLCF/Tools/fixrec.ML +++ b/src/HOL/HOLCF/Tools/fixrec.ML @@ -1,403 +1,403 @@ (* Title: HOL/HOLCF/Tools/fixrec.ML Author: Amber Telfer and Brian Huffman Recursive function definition package for HOLCF. *) signature FIXREC = sig val add_fixrec: (binding * typ option * mixfix) list -> (bool * (Attrib.binding * term)) list -> local_theory -> local_theory val add_fixrec_cmd: (binding * string option * mixfix) list -> (bool * (Attrib.binding * string)) list -> local_theory -> local_theory val add_matchers: (string * string) list -> theory -> theory val fixrec_simp_tac: Proof.context -> int -> tactic end structure Fixrec : FIXREC = struct open HOLCF_Library infixr 6 ->> infix -->> infix 9 ` val def_cont_fix_eq = @{thm def_cont_fix_eq} val def_cont_fix_ind = @{thm def_cont_fix_ind} fun fixrec_err s = error ("fixrec definition error:\n" ^ s) (*************************************************************************) (***************************** building types ****************************) (*************************************************************************) local fun binder_cfun \<^Type>\cfun T U\ = T :: binder_cfun U | binder_cfun \<^Type>\fun T U\ = T :: binder_cfun U | binder_cfun _ = [] fun body_cfun \<^Type>\cfun _ U\ = body_cfun U | body_cfun \<^Type>\fun _ U\ = body_cfun U | body_cfun T = T in fun matcherT (T, U) = body_cfun T ->> (binder_cfun T -->> U) ->> U end (*************************************************************************) (***************************** building terms ****************************) (*************************************************************************) val mk_trp = HOLogic.mk_Trueprop (* splits a cterm into the right and lefthand sides of equality *) fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t) (* similar to Thm.head_of, but for continuous application *) fun chead_of \<^Const_>\Rep_cfun _ _ for f _\ = chead_of f | chead_of u = u infix 1 === val (op ===) = HOLogic.mk_eq fun mk_mplus (t, u) = let val T = dest_matchT (Term.fastype_of t) in \<^Const>\Fixrec.mplus T\ ` t ` u end fun mk_run t = let val mT = Term.fastype_of t val T = dest_matchT mT val run = \<^Const>\Fixrec.run T\ in case t of \<^Const_>\Rep_cfun _ _\ $ \<^Const_>\Fixrec.succeed _\ $ u => u | _ => run ` t end (*************************************************************************) (************* fixed-point definitions and unfolding theorems ************) (*************************************************************************) structure FixrecUnfoldData = Generic_Data ( type T = thm Symtab.table val empty = Symtab.empty fun merge data : T = Symtab.merge (K true) data ) local fun name_of (Const (n, _)) = n | name_of (Free (n, _)) = n | name_of t = raise TERM ("Fixrec.add_unfold: lhs not a constant", [t]) val lhs_name = name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of in val add_unfold : attribute = Thm.declaration_attribute (fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th))) end fun add_fixdefs (fixes : ((binding * typ) * mixfix) list) (spec : (Attrib.binding * term) list) (lthy : local_theory) = let val thy = Proof_Context.theory_of lthy val names = map (Binding.name_of o fst o fst) fixes val all_names = space_implode "_" names val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec) val functional = lambda_tuple lhss (mk_tuple rhss) val fixpoint = mk_fix (mk_cabs functional) val cont_thm = let val prop = mk_trp (mk_cont functional) fun err _ = error ( "Continuity proof failed please check that cont2cont rules\n" ^ "or simp rules are configured for all non-HOLCF constants.\n" ^ "The error occurred for the goal statement:\n" ^ Syntax.string_of_term lthy prop) val rules = Named_Theorems.get lthy \<^named_theorems>\cont2cont\ val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac lthy (rev rules))) val slow_tac = SOLVED' (simp_tac lthy) val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err in Goal.prove lthy [] [] prop (K tac) end fun one_def (Free(n,_)) r = let val b = Long_Name.base_name n in ((Binding.name (Thm.def_name b), []), r) end | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form" fun defs [] _ = [] | defs (l::[]) r = [one_def l r] | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r) val fixdefs = defs lhss fixpoint val (fixdef_thms : (term * (string * thm)) list, lthy) = lthy |> fold_map Local_Theory.define (map (apfst fst) fixes ~~ fixdefs) fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2] val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms) val P = Var (("P", 0), map Term.fastype_of lhss ---> \<^Type>\bool\) val predicate = lambda_tuple lhss (list_comb (P, lhss)) val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm]) |> Thm.instantiate' [] [SOME (Thm.global_cterm_of thy predicate)] |> Local_Defs.unfold lthy @{thms split_paired_all split_conv split_strict} val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) |> Local_Defs.unfold lthy @{thms split_conv} fun unfolds [] _ = [] | unfolds (n::[]) thm = [(n, thm)] | unfolds (n::ns) thm = let val thmL = thm RS @{thm Pair_eqD1} val thmR = thm RS @{thm Pair_eqD2} in (n, thmL) :: unfolds ns thmR end val unfold_thms = unfolds names tuple_unfold_thm val induct_note : Attrib.binding * Thm.thm list = let val thm_name = Binding.qualify true all_names (Binding.name "induct") in ((thm_name, []), [tuple_induct_thm]) end fun unfold_note (name, thm) : Attrib.binding * Thm.thm list = let val thm_name = Binding.qualify true name (Binding.name "unfold") - val src = Attrib.internal (K add_unfold) + val src = Attrib.internal \<^here> (K add_unfold) in ((thm_name, [src]), [thm]) end val (_, lthy) = lthy |> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms) in (lthy, names, fixdef_thms, map snd unfold_thms) end (*************************************************************************) (*********** monadic notation and pattern matching compilation ***********) (*************************************************************************) structure FixrecMatchData = Theory_Data ( type T = string Symtab.table val empty = Symtab.empty fun merge data = Symtab.merge (K true) data ) (* associate match functions with pattern constants *) fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms) fun taken_names (t : term) : bstring list = let fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | taken (Free(a,_) , bs) = insert (op =) a bs | taken (f $ u , bs) = taken (f, taken (u, bs)) | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs) | taken (_ , bs) = bs in taken (t, []) end (* builds a monadic term for matching a pattern *) (* returns (rhs, free variable, used varnames) *) fun compile_pat match_name pat rhs taken = let fun comp_pat p rhs taken = if is_Free p then (rhs, p, taken) else comp_con (fastype_of p) p rhs [] taken (* compiles a monadic term for a constructor pattern *) and comp_con T p rhs vs taken = case p of \<^Const_>\Rep_cfun _ _ for f x\ => let val (rhs', v, taken') = comp_pat x rhs taken in comp_con T f rhs' (v::vs) taken' end | f $ x => let val (rhs', v, taken') = comp_pat x rhs taken in comp_con T f rhs' (v::vs) taken' end | Const (c, cT) => let val n = singleton (Name.variant_list taken) "v" val v = Free(n, T) val m = Const(match_name c, matcherT (cT, fastype_of rhs)) val k = big_lambdas vs rhs in (m`v`k, v, n::taken) end | _ => raise TERM ("fixrec: invalid pattern ", [p]) in comp_pat pat rhs taken end (* builds a monadic term for matching a function definition pattern *) (* returns (constant, (vars, matcher)) *) fun compile_lhs match_name pat rhs vs taken = case pat of \<^Const_>\Rep_cfun _ _ for f x\ => let val (rhs', v, taken') = compile_pat match_name x rhs taken in compile_lhs match_name f rhs' (v::vs) taken' end | Free(_,_) => (pat, (vs, rhs)) | Const(_,_) => (pat, (vs, rhs)) | _ => fixrec_err ("invalid function pattern: " ^ ML_Syntax.print_term pat) fun strip_alls t = (case try Logic.dest_all_global t of SOME (_, u) => strip_alls u | NONE => t) fun compile_eq match_name eq = let val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq)) in compile_lhs match_name lhs (mk_succeed rhs) [] (taken_names eq) end (* this is the pattern-matching compiler function *) fun compile_eqs match_name eqs = let val (consts, matchers) = ListPair.unzip (map (compile_eq match_name) eqs) val const = case distinct (op =) consts of [n] => n | [] => fixrec_err "no defining equations for function" | _ => fixrec_err "all equations in block must define the same function" val vars = case distinct (op = o apply2 length) (map fst matchers) of [vars] => vars | _ => fixrec_err "all equations in block must have the same arity" (* rename so all matchers use same free variables *) fun rename (vs, t) = Term.subst_free (filter_out (op =) (vs ~~ vars)) t val rhs = big_lambdas vars (mk_run (foldr1 mk_mplus (map rename matchers))) in mk_trp (const === rhs) end (*************************************************************************) (********************** Proving associated theorems **********************) (*************************************************************************) fun eta_tac i = CONVERSION Thm.eta_conversion i fun fixrec_simp_tac ctxt = let val tab = FixrecUnfoldData.get (Context.Proof ctxt) val concl = HOLogic.dest_Trueprop o Logic.strip_imp_concl o strip_alls fun tac (t, i) = let val (c, _) = (dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t val unfold_thm = the (Symtab.lookup tab c) val rule = unfold_thm RS @{thm ssubst_lhs} in CHANGED (resolve_tac ctxt [rule] i THEN eta_tac i THEN asm_simp_tac ctxt i) end in SUBGOAL (fn ti => the_default no_tac (try tac ti)) end (* proves a block of pattern matching equations as theorems, using unfold *) fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = let val rule = unfold_thm RS @{thm ssubst_lhs} val tac = resolve_tac ctxt [rule] 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1 fun prove_term t = Goal.prove ctxt [] [] t (K tac) fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t) in map prove_eqn eqns end (*************************************************************************) (************************* Main fixrec function **************************) (*************************************************************************) local (* code adapted from HOL/Tools/Datatype/primrec.ML *) fun gen_fixrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) (raw_spec' : (bool * (Attrib.binding * 'b)) list) (lthy : local_theory) = let val (skips, raw_spec) = ListPair.unzip raw_spec' val (fixes : ((binding * typ) * mixfix) list, spec : (Attrib.binding * term) list) = fst (prep_spec raw_fixes (map (fn s => (s, [], [])) raw_spec) lthy) val names = map (Binding.name_of o fst o fst) fixes fun check_head name = member (op =) names name orelse fixrec_err ("Illegal equation head. Expected " ^ commas_quote names) val chead_of_spec = chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd fun name_of (Free (n, _)) = tap check_head n | name_of _ = fixrec_err ("unknown term") val all_names = map (name_of o chead_of_spec) spec fun block_of_name n = map_filter (fn (m,eq) => if m = n then SOME eq else NONE) (all_names ~~ (spec ~~ skips)) val blocks = map block_of_name names val matcher_tab = FixrecMatchData.get (Proof_Context.theory_of lthy) fun match_name c = case Symtab.lookup matcher_tab c of SOME m => m | NONE => fixrec_err ("unknown pattern constructor: " ^ c) val matches = map (compile_eqs match_name) (map (map (snd o fst)) blocks) val spec' = map (pair Binding.empty_atts) matches val (lthy, _, _, unfold_thms) = add_fixdefs fixes spec' lthy val blocks' = map (map fst o filter_out snd) blocks val simps : (Attrib.binding * thm) list list = map (make_simps lthy) (unfold_thms ~~ blocks') fun mk_bind n : Attrib.binding = (Binding.qualify true n (Binding.name "simps"), @{attributes [simp]}) val simps1 : (Attrib.binding * thm list) list = map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps) val simps2 : (Attrib.binding * thm list) list = map (apsnd (fn thm => [thm])) (flat simps) val (_, lthy) = lthy |> fold_map Local_Theory.note (simps1 @ simps2) in lthy end in val add_fixrec = gen_fixrec Specification.check_multi_specs val add_fixrec_cmd = gen_fixrec Specification.read_multi_specs end (* local *) (*************************************************************************) (******************************** Parsers ********************************) (*************************************************************************) val opt_thm_name' : (bool * Attrib.binding) parser = \<^keyword>\(\ -- \<^keyword>\unchecked\ -- \<^keyword>\)\ >> K (true, Binding.empty_atts) || Parse_Spec.opt_thm_name ":" >> pair false val spec' : (bool * (Attrib.binding * string)) parser = opt_thm_name' -- Parse.prop >> (fn ((a, b), c) => (a, (b, c))) val multi_specs' : (bool * (Attrib.binding * string)) list parser = let val unexpected = Scan.ahead (Parse.name || \<^keyword>\[\ || \<^keyword>\(\) in Parse.enum1 "|" (spec' --| Scan.option (unexpected -- Parse.!!! \<^keyword>\|\)) end val _ = Outer_Syntax.local_theory \<^command_keyword>\fixrec\ "define recursive functions (HOLCF)" (Parse.vars -- (Parse.where_ |-- Parse.!!! multi_specs') >> (fn (vars, specs) => add_fixrec_cmd vars specs)) end diff --git a/src/HOL/Library/adhoc_overloading.ML b/src/HOL/Library/adhoc_overloading.ML --- a/src/HOL/Library/adhoc_overloading.ML +++ b/src/HOL/Library/adhoc_overloading.ML @@ -1,245 +1,245 @@ (* Author: Alexander Krauss, TU Muenchen Author: Christian Sternagel, University of Innsbruck Adhoc overloading of constants based on their types. *) signature ADHOC_OVERLOADING = sig val is_overloaded: Proof.context -> string -> bool val generic_add_overloaded: string -> Context.generic -> Context.generic val generic_remove_overloaded: string -> Context.generic -> Context.generic val generic_add_variant: string -> term -> Context.generic -> Context.generic (*If the list of variants is empty at the end of "generic_remove_variant", then "generic_remove_overloaded" is called implicitly.*) val generic_remove_variant: string -> term -> Context.generic -> Context.generic val show_variants: bool Config.T end structure Adhoc_Overloading: ADHOC_OVERLOADING = struct val show_variants = Attrib.setup_config_bool \<^binding>\show_variants\ (K false); (* errors *) fun err_duplicate_variant oconst = error ("Duplicate variant of " ^ quote oconst); fun err_not_a_variant oconst = error ("Not a variant of " ^ quote oconst); fun err_not_overloaded oconst = error ("Constant " ^ quote oconst ^ " is not declared as overloaded"); fun err_unresolved_overloading ctxt0 (c, T) t instances = let val ctxt = Config.put show_variants true ctxt0 val const_space = Proof_Context.const_space ctxt val prt_const = Pretty.block [Name_Space.pretty ctxt const_space c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)] in error (Pretty.string_of (Pretty.chunks [ Pretty.block [ Pretty.str "Unresolved adhoc overloading of constant", Pretty.brk 1, prt_const, Pretty.brk 1, Pretty.str "in term", Pretty.brk 1, Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t)]], Pretty.block ( (if null instances then [Pretty.str "no instances"] else Pretty.fbreaks ( Pretty.str "multiple instances:" :: map (Pretty.mark Markup.item o Syntax.pretty_term ctxt) instances)))])) end; (* generic data *) fun variants_eq ((v1, T1), (v2, T2)) = Term.aconv_untyped (v1, v2) andalso T1 = T2; structure Overload_Data = Generic_Data ( type T = {variants : (term * typ) list Symtab.table, oconsts : string Termtab.table}; val empty = {variants = Symtab.empty, oconsts = Termtab.empty}; fun merge ({variants = vtab1, oconsts = otab1}, {variants = vtab2, oconsts = otab2}) : T = let fun merge_oconsts _ (oconst1, oconst2) = if oconst1 = oconst2 then oconst1 else err_duplicate_variant oconst1; in {variants = Symtab.merge_list variants_eq (vtab1, vtab2), oconsts = Termtab.join merge_oconsts (otab1, otab2)} end; ); fun map_tables f g = Overload_Data.map (fn {variants = vtab, oconsts = otab} => {variants = f vtab, oconsts = g otab}); val is_overloaded = Symtab.defined o #variants o Overload_Data.get o Context.Proof; val get_variants = Symtab.lookup o #variants o Overload_Data.get o Context.Proof; val get_overloaded = Termtab.lookup o #oconsts o Overload_Data.get o Context.Proof; fun generic_add_overloaded oconst context = if is_overloaded (Context.proof_of context) oconst then context else map_tables (Symtab.update (oconst, [])) I context; fun generic_remove_overloaded oconst context = let fun remove_oconst_and_variants context oconst = let val remove_variants = (case get_variants (Context.proof_of context) oconst of NONE => I | SOME vs => fold (Termtab.remove (op =) o rpair oconst o fst) vs); in map_tables (Symtab.delete_safe oconst) remove_variants context end; in if is_overloaded (Context.proof_of context) oconst then remove_oconst_and_variants context oconst else err_not_overloaded oconst end; local fun generic_variant add oconst t context = let val ctxt = Context.proof_of context; val _ = if is_overloaded ctxt oconst then () else err_not_overloaded oconst; val T = t |> fastype_of; val t' = Term.map_types (K dummyT) t; in if add then let val _ = (case get_overloaded ctxt t' of NONE => () | SOME oconst' => err_duplicate_variant oconst'); in map_tables (Symtab.cons_list (oconst, (t', T))) (Termtab.update (t', oconst)) context end else let val _ = if member variants_eq (the (get_variants ctxt oconst)) (t', T) then () else err_not_a_variant oconst; in map_tables (Symtab.map_entry oconst (remove1 variants_eq (t', T))) (Termtab.delete_safe t') context |> (fn context => (case get_variants (Context.proof_of context) oconst of SOME [] => generic_remove_overloaded oconst context | _ => context)) end end; in val generic_add_variant = generic_variant true; val generic_remove_variant = generic_variant false; end; (* check / uncheck *) fun unifiable_with thy T1 T2 = let val maxidx1 = Term.maxidx_of_typ T1; val T2' = Logic.incr_tvar (maxidx1 + 1) T2; val maxidx2 = Term.maxidx_typ T2' maxidx1; in can (Sign.typ_unify thy (T1, T2')) (Vartab.empty, maxidx2) end; fun get_candidates ctxt (c, T) = get_variants ctxt c |> Option.map (map_filter (fn (t, T') => if unifiable_with (Proof_Context.theory_of ctxt) T T' then SOME t else NONE)); fun insert_variants ctxt t (oconst as Const (c, T)) = (case get_candidates ctxt (c, T) of SOME [] => err_unresolved_overloading ctxt (c, T) t [] | SOME [variant] => variant | _ => oconst) | insert_variants _ _ oconst = oconst; fun insert_overloaded ctxt = let fun proc t = Term.map_types (K dummyT) t |> get_overloaded ctxt |> Option.map (Const o rpair (Term.type_of t)); in Pattern.rewrite_term_top (Proof_Context.theory_of ctxt) [] [proc] end; fun check ctxt = map (fn t => Term.map_aterms (insert_variants ctxt t) t); fun uncheck ctxt ts = if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts else map (insert_overloaded ctxt) ts; fun reject_unresolved ctxt = let val the_candidates = the o get_candidates ctxt; fun check_unresolved t = (case filter (is_overloaded ctxt o fst) (Term.add_consts t []) of [] => t | (cT :: _) => err_unresolved_overloading ctxt cT t (the_candidates cT)); in map check_unresolved end; (* setup *) val _ = Context.>> (Syntax_Phases.term_check 0 "adhoc_overloading" check #> Syntax_Phases.term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved #> Syntax_Phases.term_uncheck 0 "adhoc_overloading" uncheck); (* commands *) fun generic_adhoc_overloading_cmd add = if add then fold (fn (oconst, ts) => generic_add_overloaded oconst #> fold (generic_add_variant oconst) ts) else fold (fn (oconst, ts) => fold (generic_remove_variant oconst) ts); fun adhoc_overloading_cmd' add args phi = let val args' = args |> map (apsnd (map_filter (fn t => let val t' = Morphism.term phi t; in if Term.aconv_untyped (t, t') then SOME t' else NONE end))); in generic_adhoc_overloading_cmd add args' end; fun adhoc_overloading_cmd add raw_args lthy = let fun const_name ctxt = fst o dest_Const o Proof_Context.read_const {proper = false, strict = false} ctxt; (* FIXME {proper = true, strict = true} (!?) *) fun read_term ctxt = singleton (Variable.polymorphic ctxt) o Syntax.read_term ctxt; val args = raw_args |> map (apfst (const_name lthy)) |> map (apsnd (map (read_term lthy))); in - Local_Theory.declaration {syntax = true, pervasive = false} + Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (adhoc_overloading_cmd' add args) lthy end; val _ = Outer_Syntax.local_theory \<^command_keyword>\adhoc_overloading\ "add adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true); val _ = Outer_Syntax.local_theory \<^command_keyword>\no_adhoc_overloading\ "delete adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false); end; diff --git a/src/HOL/Nominal/nominal_inductive.ML b/src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML +++ b/src/HOL/Nominal/nominal_inductive.ML @@ -1,698 +1,698 @@ (* Title: HOL/Nominal/nominal_inductive.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. *) signature NOMINAL_INDUCTIVE = sig val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state val prove_eqvt: string -> string list -> local_theory -> local_theory end structure NominalInductive : NOMINAL_INDUCTIVE = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val fresh_prod = @{thm fresh_prod}; val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else insert (op aconv o apply2 fst) (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} => EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)), REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion; fun first_order_matchs pats objs = Thm.first_order_match (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); fun first_order_mrs ths th = ths MRS Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th; fun prove_strong_ind s avoids lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val monos = Inductive.get_monos lthy; val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP lthy (hd names))))); val (raw_induct', ctxt') = lthy |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse error ("Duplicate variable names for case " ^ quote a)); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); val avoids' = if null induct_cases then replicate (length intrs) ("", []) else map (fn name => (name, the_default [] (AList.lookup op = avoids name))) induct_cases; fun mk_avoids params (name, ps) = let val k = length params - 1 in map (fn x => case find_index (equal x o fst) params of ~1 => error ("No such variable in case " ^ quote name ^ " of inductive definition: " ^ quote x) | i => (Bound (k - i), snd (nth params i))) ps end; val prems = map (fn (prem, avoid) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, fold (add_binders thy 0) (prems @ [concl]) [] @ map (apfst (incr_boundvars 1)) (mk_avoids params avoid), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); val atomTs = distinct op = (maps (map snd o #2) prems); val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy ("fs_" ^ Long_Name.base_name (fst (dest_Type T)))) atomTs)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_distinct [] = [] | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) => if T = U then SOME (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) else NONE) xs @ mk_distinct xs; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh bvars @ mk_distinct bvars @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') in (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem))) end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val vc_compat = map (fn (params, bvars, prems, (p, ts)) => map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (mk_distinct bvars @ maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_const U T $ u $ t)) bvars) (ts ~~ binder_types (fastype_of p)))) prems; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn aT => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs; fun eqvt_ss ctxt = put_simpset HOL_basic_ss ctxt addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; val perm_bij = Global_Theory.get_thms thy "perm_bij"; val fs_atoms = map (fn aT => Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "1")) atomTs; val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'"; val fresh_atm = Global_Theory.get_thms thy "fresh_atm"; val swap_simps = Global_Theory.get_thms thy "swap_simps"; val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh"; fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = let (** protect terms to avoid that fresh_prod interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, NominalDatatype.fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn {context = goal_ctxt, ...} => EVERY [resolve_tac goal_ctxt exists_fresh' 1, resolve_tac goal_ctxt fs_atoms 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn goal_ctxt => EVERY [eresolve_tac goal_ctxt [exE] 1, full_simp_tac (put_simpset HOL_ss goal_ctxt addsimps (fresh_prod :: fresh_atm)) 1, full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps [@{thm id_apply}]) 1, REPEAT (eresolve_tac goal_ctxt [conjE] 1)]) [ex] ctxt in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; fun mk_ind_proof ctxt thss = Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} => let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} => resolve_tac goal_ctxt1 [raw_induct] 1 THEN EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => [REPEAT (resolve_tac goal_ctxt1 [allI] 1), simp_tac (eqvt_ss goal_ctxt1) 1, SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} => let val (params', (pis, z)) = chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||> split_last; val bvars' = map (fn (Bound i, T) => (nth params' (length params' - i), T) | (t, T) => (t, T)) bvars; val pi_bvars = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val (freshs1, freshs2, ctxt') = fold (obtain_fresh_name (ts @ pi_bvars)) (map snd bvars') ([], [], goal_ctxt2); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold (concat_perm #> map) pis' pis; val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) (Vartab.empty, Vartab.empty); val ihyp' = Thm.instantiate (TVars.empty, Vars.make (map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t)) (map (Envir.subst_term env) vs ~~ map (fold_rev (NominalDatatype.mk_perm []) (rev pis' @ pis)) params' @ [z]))) ihyp; fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (eqvt_ss ctxt') (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt') (rev pis' @ pis) th)); val (gprems1, gprems2) = split_list (map (fn (th, t) => if null (preds_of ps t) then (SOME th, mk_pi th) else (map_thm ctxt' (split_conj (K o I) names) (eresolve_tac ctxt' [conjunct1] 1) monos NONE th, mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th)))) (gprems ~~ oprems)) |>> map_filter I; val vc_compat_ths' = map (fn th => let val th' = first_order_mrs gprems1 th; val (bop, lhs, rhs) = (case Thm.concl_of th' of _ $ (fresh $ lhs $ rhs) => (fn t => fn u => fresh $ t $ u, lhs, rhs) | _ $ (_ $ (_ $ lhs $ rhs)) => (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); val th'' = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) (fold_rev (NominalDatatype.mk_perm []) pis rhs))) (fn {context = goal_ctxt3, ...} => simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps (fresh_bij @ perm_bij)) 1 THEN resolve_tac goal_ctxt3 [th'] 1) in Simplifier.simplify (eqvt_ss ctxt' addsimps fresh_atm) th'' end) vc_compat_ths; val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) (** we have to pre-simplify the rewrite rules **) val swap_simps_simpset = put_simpset HOL_ss ctxt' addsimps swap_simps @ map (Simplifier.simplify (put_simpset HOL_ss ctxt' addsimps swap_simps)) (vc_compat_ths'' @ freshs2'); val th = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn {context = goal_ctxt4, ...} => EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1, resolve_tac goal_ctxt4 [ihyp'] 1, REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems) (simp_tac swap_simps_simpset 1), REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss goal_ctxt4 addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac goal_ctxt4 gprems2 1)])); val final = Goal.prove ctxt' [] [] (Thm.term_of concl) (fn {context = goal_ctxt5, ...} => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5 addsimps vc_compat_ths'' @ freshs2' @ perm_fresh_fresh @ fresh_atm) 1); val final' = Proof_Context.export ctxt' goal_ctxt2 [final]; in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1]) (prems ~~ thss ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN asm_full_simp_tac goal_ctxt 1) end) |> singleton (Proof_Context.export ctxt lthy); (** strong case analysis rule **) val cases_prems = map (fn ((name, avoids), rule) => let val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] lthy; val prem :: prems = Logic.strip_imp_prems rule'; val concl = Logic.strip_imp_concl rule' in (prem, List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), concl, fold_map (fn (prem, (_, avoid)) => fn ctxt => let val prems = Logic.strip_assums_hyp prem; val params = Logic.strip_params prem; val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) = if member (op = o apsnd fst) bnds (Bound i) then let val ([s'], ctxt') = Variable.variant_fixes [s] ctxt; val t = Free (s', T) in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts); val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params (0, 0, ctxt, [], [], [], []) in ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt') end) (prems ~~ avoids) ctxt') end) (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~ elims); val cases_prems' = map (fn (prem, args, concl, (prems, _)) => let fun mk_prem (ps, [], _, prems) = Logic.list_all (ps, Logic.list_implies (prems, concl)) | mk_prem (ps, qs, _, prems) = Logic.list_all (ps, Logic.mk_implies (Logic.list_implies (mk_distinct qs @ maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) args) qs, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop prems))), concl)) in map mk_prem prems end) cases_prems; fun cases_eqvt_simpset ctxt = put_simpset HOL_ss ctxt addsimps eqvt_thms @ swap_simps @ perm_pi_simp addsimprocs [NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; fun simp_fresh_atm ctxt = Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps fresh_atm); fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt))), prems') = (name, Goal.prove ctxt [] (prem :: prems') concl (fn {prems = hyp :: hyps, context = ctxt1} => EVERY (resolve_tac ctxt1 [hyp RS elim] 1 :: map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} => if null qs then resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1 else let val params' = map (Thm.term_of o #2 o nth (rev params)) is; val tab = params' ~~ map fst qs; val (hyps1, hyps2) = chop (length args) case_hyps; (* turns a = t and [x1 # t, ..., xn # t] *) (* into [x1 # a, ..., xn # a] *) fun inst_fresh th' ths = let val (ths1, ths2) = chop (length qs) ths in (map (fn th => let val (cf, ct) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)); val arg_cong' = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ct)] [NONE, SOME ct, SOME cf] (arg_cong RS iffD2); val inst = Thm.first_order_match (ct, Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th'))) in [th', th] MRS Thm.instantiate inst arg_cong' end) ths1, ths2) end; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length args * length qs) (map (first_order_mrs hyps2) vc_compat_ths); val vc_compat_ths' = NominalDatatype.mk_not_sym vc_compat_ths1 @ flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); val (freshs1, freshs2, ctxt3) = fold (obtain_fresh_name (args @ map fst qs @ params')) (map snd qs) ([], [], ctxt2); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis = map (NominalDatatype.perm_of_pair) ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); val mk_pis = fold_rev (mk_perm_bool ctxt3) (map (Thm.cterm_of ctxt3) pis); val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms (fn x as Free _ => if member (op =) args x then x else (case AList.lookup op = tab x of SOME y => y | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps)); val inst = Thm.first_order_match (Thm.dest_arg (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); val th = Goal.prove ctxt3 [] [] (Thm.term_of concl) (fn {context = ctxt4, ...} => resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} => let val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; val case_simpset = cases_eqvt_simpset ctxt5 addsimps freshs2' @ map (simp_fresh_atm ctxt5) (vc_compat_ths' @ fresh_hyps'); val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; val hyps1' = map (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; val hyps2' = map (mk_pis #> Simplifier.simplify case_simpset) hyps2; val case_hyps' = hyps1' @ hyps2' in simp_tac case_simpset 1 THEN REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN resolve_tac ctxt5 case_hyps' 1) end) ctxt4 1) val final = Proof_Context.export ctxt3 ctxt2 [th] in resolve_tac ctxt2 final 1 end) ctxt1 1) (thss ~~ hyps ~~ prems))) |> singleton (Proof_Context.export ctxt lthy)) in ctxt'' |> Proof.theorem NONE (fn thss => fn lthy1 => let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr lthy1)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1; val strong_cases = map (mk_cases_proof ##> Inductive.rulify lthy1) (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); val strong_induct_atts = - map (Attrib.internal o K) + map (Attrib.internal \<^here> o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects lthy1 (1 upto length names) strong_induct'; in lthy2 |> Local_Theory.notes [((rec_qualified (Binding.name "strong_inducts"), []), strong_inducts |> map (fn th => ([th], - [Attrib.internal (K ind_case_names), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> + [Attrib.internal \<^here> (K ind_case_names), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> Local_Theory.notes (map (fn ((name, elim), (_, cases)) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), - [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) + [Attrib.internal \<^here> (K (Rule_Cases.case_names (map snd cases))), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) (strong_cases ~~ induct_cases')) |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; fun prove_eqvt s xatoms lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val intrs = map (atomize_intr lthy) intrs; val monos = Inductive.get_monos lthy; val intrs' = Inductive.unpartition_rules intrs (map (fn (((s, ths), (_, k)), th) => (s, ths ~~ Inductive.infer_intro_vars thy th k ths)) (Inductive.partition_rules raw_induct intrs ~~ Inductive.arities_of raw_induct ~~ elims)); val k = length (Inductive.params_of raw_induct); val atoms' = NominalAtoms.atoms_of thy; val atoms = if null xatoms then atoms' else let val atoms = map (Sign.intern_type thy) xatoms in (case duplicates op = atoms of [] => () | xs => error ("Duplicate atoms: " ^ commas xs); case subtract (op =) atoms' atoms of [] => () | xs => error ("No such atoms: " ^ commas xs); atoms) end; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val (([t], [pi]), ctxt1) = lthy |> Variable.import_terms false [Thm.concl_of raw_induct] ||>> Variable.variant_fixes ["pi"]; fun eqvt_simpset ctxt = put_simpset HOL_basic_ss ctxt addsimps (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs [mk_perm_bool_simproc names, NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val ps = map (fst o HOLogic.dest_imp) (HOLogic.dest_conj (HOLogic.dest_Trueprop t)); fun eqvt_tac ctxt pi (intr, vs) st = let fun eqvt_err s = let val ([t], ctxt') = Variable.import_terms true [Thm.prop_of intr] ctxt in error ("Could not prove equivariance for introduction rule\n" ^ Syntax.string_of_term ctxt' t ^ "\n" ^ s) end; val res = SUBPROOF (fn {context = goal_ctxt, prems, params, ...} => let val prems' = map (fn th => the_default th (map_thm goal_ctxt (split_conj (K I) names) (eresolve_tac goal_ctxt [conjunct2] 1) monos NONE th)) prems; val prems'' = map (fn th => Simplifier.simplify (eqvt_simpset goal_ctxt) (mk_perm_bool goal_ctxt (Thm.cterm_of goal_ctxt pi) th)) prems'; val intr' = infer_instantiate goal_ctxt (map (#1 o dest_Var) vs ~~ map (Thm.cterm_of goal_ctxt o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params) intr in (resolve_tac goal_ctxt [intr'] THEN_ALL_NEW (TRY o resolve_tac goal_ctxt prems'')) 1 end) ctxt 1 st in case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of NONE => eqvt_err ("Rule does not match goal\n" ^ Syntax.string_of_term ctxt (hd (Thm.prems_of st))) | SOME (th, _) => Seq.single th end; 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)) (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt1 [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => let val (h, ts) = strip_comb p; val (ts1, ts2) = chop k ts in HOLogic.mk_imp (p, list_comb (h, ts1 @ map (NominalDatatype.mk_perm [] pi') ts2)) end) ps))) (fn {context = goal_ctxt, ...} => EVERY (resolve_tac goal_ctxt [raw_induct] 1 :: map (fn intr_vs => full_simp_tac (eqvt_simpset goal_ctxt) 1 THEN eqvt_tac goal_ctxt pi' intr_vs) intrs')) |> singleton (Proof_Context.export ctxt1 lthy))) end) atoms in lthy |> Local_Theory.notes (map (fn (name, ths) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), - [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) + [Attrib.internal \<^here> (K NominalThmDecls.eqvt_add)]), [(ths, [])])) (names ~~ transp thss)) |> snd end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive\ "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\avoids\ |-- Parse.and_list1 (Parse.name -- (\<^keyword>\:\ |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) => prove_strong_ind name avoids)); val _ = Outer_Syntax.local_theory \<^command_keyword>\equivariance\ "prove equivariance for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\[\ |-- Parse.list1 Parse.name --| \<^keyword>\]\) [] >> (fn (name, atoms) => prove_eqvt name atoms)); end diff --git a/src/HOL/Nominal/nominal_inductive2.ML b/src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML +++ b/src/HOL/Nominal/nominal_inductive2.ML @@ -1,505 +1,505 @@ (* Title: HOL/Nominal/nominal_inductive2.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. Experimental version that allows to avoid lists of atoms. *) signature NOMINAL_INDUCTIVE2 = sig val prove_strong_ind: string -> string option -> (string * string list) list -> local_theory -> Proof.state end structure NominalInductive2 : NOMINAL_INDUCTIVE2 = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)); fun fresh_postprocess ctxt = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim}, @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else AList.map_default op = (T, []) (insert op aconv (incr_boundvars (~i) u))) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} => EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)), REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; fun abs_params params t = let val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params) in (Logic.list_all (params, t), (rev vs, subst_bounds (vs, t))) end; fun inst_params thy (vs, p) th cts = let val env = Pattern.first_order_match thy (p, Thm.prop_of th) (Vartab.empty, Vartab.empty) in Thm.instantiate (TVars.empty, Vars.make (map (dest_Var o Envir.subst_term env) vs ~~ cts)) th end; fun prove_strong_ind s alt_name avoids lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val monos = Inductive.get_monos lthy; val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP lthy (hd names))))); val induct_cases' = if null induct_cases then replicate (length intrs) "" else induct_cases; val (raw_induct', ctxt') = lthy |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); fun mk_avoids params name sets = let val (_, ctxt') = Proof_Context.add_fixes (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) lthy; fun mk s = let val t = Syntax.read_term ctxt' s; val t' = fold_rev absfree params t |> funpow (length params) (fn Abs (_, _, t) => t) in (t', HOLogic.dest_setT (fastype_of t)) end handle TERM _ => error ("Expression " ^ quote s ^ " to be avoided in case " ^ quote name ^ " is not a set type"); fun add_set p [] = [p] | add_set (t, T) ((u, U) :: ps) = if T = U then let val S = HOLogic.mk_setT T in (Const (\<^const_name>\sup\, S --> S --> S) $ u $ t, T) :: ps end else (u, U) :: add_set (t, T) ps in fold (mk #> add_set) sets [] end; val prems = map (fn (prem, name) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, if null avoids then map (fn (T, ts) => (HOLogic.mk_set T ts, T)) (fold (add_binders thy 0) (prems @ [concl]) []) else case AList.lookup op = avoids name of NONE => [] | SOME sets => map (apfst (incr_boundvars 1)) (mk_avoids params name sets), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ induct_cases'); val atomTs = distinct op = (maps (map snd o #2) prems); val atoms = map (fst o dest_Type) atomTs; val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn a => Sign.intern_class thy ("fs_" ^ Long_Name.base_name a)) atoms)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh sets @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); in abs_params params' prem end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) => map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const U T $ u $ t)) sets) (ts ~~ binder_types (fastype_of p)) @ map (fn (u, U) => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |> split_list) prems |> split_list; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn a => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name a ^ "2")) atoms; fun eqvt_ss ctxt = put_simpset HOL_basic_ss ctxt addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij"; val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms; val at_insts = map (NominalAtoms.at_inst_of thy) atoms; val dj_thms = maps (fn a => map (NominalAtoms.dj_thm_of thy a) (remove (op =) a atoms)) atoms; val finite_ineq = map2 (fn th => fn th' => th' RS (th RS @{thm pt_set_finite_ineq})) pt_insts at_insts; val perm_set_forget = map (fn th => th RS @{thm dj_perm_set_forget}) dj_thms; val perm_freshs_freshs = atomTs ~~ map2 (fn th => fn th' => th' RS (th RS @{thm pt_freshs_freshs})) pt_insts at_insts; fun obtain_fresh_name ts sets (T, fin) (freshs, ths1, ths2, ths3, ctxt) = let val thy = Proof_Context.theory_of ctxt; (** protect terms to avoid that fresh_star_prod_set interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts); val atom = fst (dest_Type T); val {at_inst, ...} = NominalAtoms.the_atom_info thy atom; val fs_atom = Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name atom ^ "1"); val avoid_th = Thm.instantiate' [SOME (Thm.global_ctyp_of thy (fastype_of p))] [SOME (Thm.global_cterm_of thy p)] ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding}); val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result (fn ctxt' => EVERY [resolve_tac ctxt' [avoid_th] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1, full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, rotate_tac 1 1, REPEAT (eresolve_tac ctxt' [conjE] 1)]) [] ctxt; val (Ts1, _ :: Ts2) = chop_prefix (not o equal T) (map snd sets); val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2); val (pis1, pis2) = chop (length Ts1) (map Bound (length pTs - 1 downto 0)); val _ $ (f $ (_ $ pi $ l) $ r) = Thm.prop_of th2 val th2' = Goal.prove ctxt' [] [] (Logic.list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop (f $ fold_rev (NominalDatatype.mk_perm (rev pTs)) (pis1 @ pi :: pis2) l $ r))) (fn {context = goal_ctxt, ...} => cut_facts_tac [th2] 1 THEN full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps perm_set_forget) 1) |> Simplifier.simplify (eqvt_ss ctxt') in (freshs @ [Thm.term_of cx], ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt') end; fun mk_ind_proof ctxt thss = Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} => let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} => resolve_tac goal_ctxt1 [raw_induct] 1 THEN EVERY (maps (fn (((((_, sets, oprems, _), vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) => [REPEAT (resolve_tac goal_ctxt1 [allI] 1), simp_tac (eqvt_ss goal_ctxt1) 1, SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} => let val (cparams', (pis, z)) = chop (length params - length atomTs - 1) (map #2 params) ||> (map Thm.term_of #> split_last); val params' = map Thm.term_of cparams' val sets' = map (apfst (curry subst_bounds (rev params'))) sets; val pi_sets = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) sets'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val gprems1 = map_filter (fn (th, t) => if null (preds_of ps t) then SOME th else map_thm goal_ctxt2 (split_conj (K o I) names) (eresolve_tac goal_ctxt2 [conjunct1] 1) monos NONE th) (gprems ~~ oprems); val vc_compat_ths' = map2 (fn th => fn p => let val th' = gprems1 MRS inst_params thy p th cparams'; val (h, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of th')) in Goal.prove goal_ctxt2 [] [] (HOLogic.mk_Trueprop (list_comb (h, map (fold_rev (NominalDatatype.mk_perm []) pis) ts))) (fn {context = goal_ctxt3, ...} => simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps (fresh_star_bij @ finite_ineq)) 1 THEN resolve_tac goal_ctxt3 [th'] 1) end) vc_compat_ths vc_compat_vs; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length sets) vc_compat_ths'; val vc_compat_ths1' = map (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.rewrite (eqvt_ss goal_ctxt2))))) vc_compat_ths1; val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold (obtain_fresh_name ts sets) (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], goal_ctxt2); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold_rev (concat_perm #> map) pis' pis; val ihyp' = inst_params thy vs_ihypt ihyp (map (fold_rev (NominalDatatype.mk_perm []) (pis' @ pis) #> Thm.global_cterm_of thy) params' @ [Thm.global_cterm_of thy z]); fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt'' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (eqvt_ss ctxt'') (fold_rev (mk_perm_bool ctxt'' o Thm.cterm_of ctxt'') (pis' @ pis) th)); val gprems2 = map (fn (th, t) => if null (preds_of ps t) then mk_pi th else mk_pi (the (map_thm ctxt'' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt'' (length pis'')) monos (SOME t) th))) (gprems ~~ oprems); val perm_freshs_freshs' = map (fn (th, (_, T)) => th RS the (AList.lookup op = perm_freshs_freshs T)) (fresh_ths2 ~~ sets); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn {context = goal_ctxt4, ...} => EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1, resolve_tac goal_ctxt4 [ihyp'] 1] @ map (fn th => resolve_tac goal_ctxt4 [th] 1) fresh_ths3 @ [REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss goal_ctxt4 addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac goal_ctxt4 gprems2 1)])); val final = Goal.prove ctxt'' [] [] (Thm.term_of concl) (fn {context = goal_ctxt5, ...} => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5 addsimps vc_compat_ths1' @ fresh_ths1 @ perm_freshs_freshs') 1); val final' = Proof_Context.export ctxt'' goal_ctxt2 [final]; in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1]) (prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN asm_full_simp_tac goal_ctxt 1) end) |> fresh_postprocess ctxt |> singleton (Proof_Context.export ctxt lthy); in ctxt'' |> Proof.theorem NONE (fn thss => fn lthy1 => let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr lthy1)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1; val strong_induct_atts = - map (Attrib.internal o K) + map (Attrib.internal \<^here> o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val (induct_name, inducts_name) = case alt_name of NONE => (rec_qualified (Binding.name "strong_induct"), rec_qualified (Binding.name "strong_inducts")) | SOME s => (Binding.name s, Binding.name (s ^ "s")); val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note ((induct_name, strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects lthy2 (1 upto length names) strong_induct' in lthy2 |> Local_Theory.notes [((inducts_name, []), strong_inducts |> map (fn th => ([th], - [Attrib.internal (K ind_case_names), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd + [Attrib.internal \<^here> (K ind_case_names), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive2\ "prove strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.option (\<^keyword>\(\ |-- Parse.!!! (Parse.name --| \<^keyword>\)\)) -- (Scan.optional (\<^keyword>\avoids\ |-- Parse.enum1 "|" (Parse.name -- (\<^keyword>\:\ |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) => prove_strong_ind name rule_name avoids)); end diff --git a/src/HOL/Real_Asymp/exp_log_expression.ML b/src/HOL/Real_Asymp/exp_log_expression.ML --- a/src/HOL/Real_Asymp/exp_log_expression.ML +++ b/src/HOL/Real_Asymp/exp_log_expression.ML @@ -1,689 +1,689 @@ signature EXP_LOG_EXPRESSION = sig exception DUP datatype expr = ConstExpr of term | X | Uminus of expr | Add of expr * expr | Minus of expr * expr | Inverse of expr | Mult of expr * expr | Div of expr * expr | Ln of expr | Exp of expr | Power of expr * term | LnPowr of expr * expr | ExpLn of expr | Powr of expr * expr | Powr_Nat of expr * expr | Powr' of expr * term | Root of expr * term | Absolute of expr | Sgn of expr | Min of expr * expr | Max of expr * expr | Floor of expr | Ceiling of expr | Frac of expr | NatMod of expr * expr | Sin of expr | Cos of expr | ArcTan of expr | Custom of string * term * expr list val preproc_term_conv : Proof.context -> conv val expr_to_term : expr -> term val reify : Proof.context -> term -> expr * thm val reify_simple : Proof.context -> term -> expr * thm type custom_handler = Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis val register_custom : binding -> term -> custom_handler -> local_theory -> local_theory val register_custom_from_thm : binding -> thm -> custom_handler -> local_theory -> local_theory val expand_custom : Proof.context -> string -> custom_handler option val to_mathematica : expr -> string val to_maple : expr -> string val to_maxima : expr -> string val to_sympy : expr -> string val to_sage : expr -> string val reify_mathematica : Proof.context -> term -> string val reify_maple : Proof.context -> term -> string val reify_maxima : Proof.context -> term -> string val reify_sympy : Proof.context -> term -> string val reify_sage : Proof.context -> term -> string val limit_mathematica : string -> string val limit_maple : string -> string val limit_maxima : string -> string val limit_sympy : string -> string val limit_sage : string -> string end structure Exp_Log_Expression : EXP_LOG_EXPRESSION = struct datatype expr = ConstExpr of term | X | Uminus of expr | Add of expr * expr | Minus of expr * expr | Inverse of expr | Mult of expr * expr | Div of expr * expr | Ln of expr | Exp of expr | Power of expr * term | LnPowr of expr * expr | ExpLn of expr | Powr of expr * expr | Powr_Nat of expr * expr | Powr' of expr * term | Root of expr * term | Absolute of expr | Sgn of expr | Min of expr * expr | Max of expr * expr | Floor of expr | Ceiling of expr | Frac of expr | NatMod of expr * expr | Sin of expr | Cos of expr | ArcTan of expr | Custom of string * term * expr list type custom_handler = Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis type entry = { name : string, pat : term, net_pat : term, expand : custom_handler } type entry' = { pat : term, net_pat : term, expand : custom_handler } exception DUP structure Custom_Funs = Generic_Data ( type T = { name_table : entry' Name_Space.table, net : entry Item_Net.T } fun eq_entry ({name = name1, ...}, {name = name2, ...}) = (name1 = name2) val empty = { name_table = Name_Space.empty_table "exp_log_custom_function", net = Item_Net.init eq_entry (fn {net_pat, ...} => [net_pat]) } fun merge ({name_table = tbl1, net = net1}, {name_table = tbl2, net = net2}) = {name_table = Name_Space.join_tables (fn _ => raise DUP) (tbl1, tbl2), net = Item_Net.merge (net1, net2)} ) fun rewrite' ctxt old_prems bounds thms ct = let val thy = Proof_Context.theory_of ctxt fun apply_rule t thm = let val lhs = thm |> Thm.concl_of |> Logic.dest_equals |> fst val _ = Pattern.first_order_match thy (lhs, t) (Vartab.empty, Vartab.empty) val insts = (lhs, t) |> apply2 (Thm.cterm_of ctxt) |> Thm.first_order_match val thm = Thm.instantiate insts thm val prems = Thm.prems_of thm val frees = fold Term.add_frees prems [] in if exists (member op = bounds o fst) frees then NONE else let val thm' = thm OF (map (Thm.assume o Thm.cterm_of ctxt) prems) val prems' = fold (insert op aconv) prems old_prems val crhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd |> Thm.cterm_of ctxt in SOME (thm', crhs, prems') end end handle Pattern.MATCH => NONE fun rewrite_subterm prems ct (Abs (x, _, _)) = let val ((v, ct'), ctxt') = Variable.dest_abs_cterm ct ctxt; val (thm, prems) = rewrite' ctxt' prems (x :: bounds) thms ct' in if Thm.is_reflexive thm then (Thm.reflexive ct, prems) else (Thm.abstract_rule x v thm, prems) end | rewrite_subterm prems ct (_ $ _) = let val (cs, ct) = Thm.dest_comb ct val (thm, prems') = rewrite' ctxt prems bounds thms cs val (thm', prems'') = rewrite' ctxt prems' bounds thms ct in (Thm.combination thm thm', prems'') end | rewrite_subterm prems ct _ = (Thm.reflexive ct, prems) val t = Thm.term_of ct in case get_first (apply_rule t) thms of NONE => rewrite_subterm old_prems ct t | SOME (thm, rhs, prems) => case rewrite' ctxt prems bounds thms rhs of (thm', prems) => (Thm.transitive thm thm', prems) end fun rewrite ctxt thms ct = let val thm1 = Thm.eta_long_conversion ct val rhs = thm1 |> Thm.cprop_of |> Thm.dest_comb |> snd val (thm2, prems) = rewrite' ctxt [] [] thms rhs val rhs = thm2 |> Thm.cprop_of |> Thm.dest_comb |> snd val thm3 = Thm.eta_conversion rhs val thm = Thm.transitive thm1 (Thm.transitive thm2 thm3) in fold (fn prem => fn thm => Thm.implies_intr (Thm.cterm_of ctxt prem) thm) prems thm end fun preproc_term_conv ctxt = let val thms = Named_Theorems.get ctxt \<^named_theorems>\real_asymp_reify_simps\ val thms = map (fn thm => thm RS @{thm HOL.eq_reflection}) thms in rewrite ctxt thms end fun register_custom' binding pat expand context = let val n = pat |> fastype_of |> strip_type |> fst |> length val maxidx = Term.maxidx_of_term pat val vars = map (fn i => Var ((Name.uu_, maxidx + i), \<^typ>\real\)) (1 upto n) val net_pat = Library.foldl betapply (pat, vars) val {name_table = tbl, net = net} = Custom_Funs.get context val entry' = {pat = pat, net_pat = net_pat, expand = expand} val (name, tbl) = Name_Space.define context true (binding, entry') tbl val entry = {name = name, pat = pat, net_pat = net_pat, expand = expand} val net = Item_Net.update entry net in Custom_Funs.put {name_table = tbl, net = net} context end fun register_custom binding pat expand = let fun decl phi = register_custom' binding (Morphism.term phi pat) expand in - Local_Theory.declaration {syntax = false, pervasive = false} decl + Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} decl end fun register_custom_from_thm binding thm expand = let val pat = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_comb |> snd in register_custom binding pat expand end fun expand_custom ctxt name = let val {name_table, ...} = Custom_Funs.get (Context.Proof ctxt) in case Name_Space.lookup name_table name of NONE => NONE | SOME {expand, ...} => SOME expand end fun expr_to_term e = let fun expr_to_term' (ConstExpr c) = c | expr_to_term' X = Bound 0 | expr_to_term' (Add (a, b)) = \<^term>\(+) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Mult (a, b)) = \<^term>\(*) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Minus (a, b)) = \<^term>\(-) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Div (a, b)) = \<^term>\(/) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Uminus a) = \<^term>\uminus :: real => _\ $ expr_to_term' a | expr_to_term' (Inverse a) = \<^term>\inverse :: real => _\ $ expr_to_term' a | expr_to_term' (Ln a) = \<^term>\ln :: real => _\ $ expr_to_term' a | expr_to_term' (Exp a) = \<^term>\exp :: real => _\ $ expr_to_term' a | expr_to_term' (Powr (a,b)) = \<^term>\(powr) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Powr_Nat (a,b)) = \<^term>\powr_nat :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (LnPowr (a,b)) = \<^term>\ln :: real => _\ $ (\<^term>\(powr) :: real => _\ $ expr_to_term' a $ expr_to_term' b) | expr_to_term' (ExpLn a) = \<^term>\exp :: real => _\ $ (\<^term>\ln :: real => _\ $ expr_to_term' a) | expr_to_term' (Powr' (a,b)) = \<^term>\(powr) :: real => _\ $ expr_to_term' a $ b | expr_to_term' (Power (a,b)) = \<^term>\(^) :: real => _\ $ expr_to_term' a $ b | expr_to_term' (Floor a) = \<^term>\Multiseries_Expansion.rfloor\ $ expr_to_term' a | expr_to_term' (Ceiling a) = \<^term>\Multiseries_Expansion.rceil\ $ expr_to_term' a | expr_to_term' (Frac a) = \<^term>\Archimedean_Field.frac :: real \ real\ $ expr_to_term' a | expr_to_term' (NatMod (a,b)) = \<^term>\Multiseries_Expansion.rnatmod\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Root (a,b)) = \<^term>\root :: nat \ real \ _\ $ b $ expr_to_term' a | expr_to_term' (Sin a) = \<^term>\sin :: real => _\ $ expr_to_term' a | expr_to_term' (ArcTan a) = \<^term>\arctan :: real => _\ $ expr_to_term' a | expr_to_term' (Cos a) = \<^term>\cos :: real => _\ $ expr_to_term' a | expr_to_term' (Absolute a) = \<^term>\abs :: real => _\ $ expr_to_term' a | expr_to_term' (Sgn a) = \<^term>\sgn :: real => _\ $ expr_to_term' a | expr_to_term' (Min (a,b)) = \<^term>\min :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Max (a,b)) = \<^term>\max :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Custom (_, t, args)) = Envir.beta_eta_contract ( fold (fn e => fn t => betapply (t, expr_to_term' e )) args t) in Abs ("x", \<^typ>\real\, expr_to_term' e) end fun reify_custom ctxt t = let val thy = Proof_Context.theory_of ctxt val t = Envir.beta_eta_contract t val t' = Envir.beta_eta_contract (Term.abs ("x", \<^typ>\real\) t) val {net, ...} = Custom_Funs.get (Context.Proof ctxt) val entries = Item_Net.retrieve_matching net (Term.subst_bound (Free ("x", \<^typ>\real\), t)) fun go {pat, name, ...} = let val n = pat |> fastype_of |> strip_type |> fst |> length val maxidx = Term.maxidx_of_term t' val vs = map (fn i => (Name.uu_, maxidx + i)) (1 upto n) val args = map (fn v => Var (v, \<^typ>\real => real\) $ Bound 0) vs val pat' = Envir.beta_eta_contract (Term.abs ("x", \<^typ>\real\) (Library.foldl betapply (pat, args))) val (T_insts, insts) = Pattern.match thy (pat', t') (Vartab.empty, Vartab.empty) fun map_option _ [] acc = SOME (rev acc) | map_option f (x :: xs) acc = case f x of NONE => NONE | SOME y => map_option f xs (y :: acc) val t = Envir.subst_term (T_insts, insts) pat in Option.map (pair (name, t)) (map_option (Option.map snd o Vartab.lookup insts) vs []) end handle Pattern.MATCH => NONE in get_first go entries end fun reify_aux ctxt t' t = let fun is_const t = fastype_of (Abs ("x", \<^typ>\real\, t)) = \<^typ>\real \ real\ andalso not (exists_subterm (fn t => t = Bound 0) t) fun is_const' t = not (exists_subterm (fn t => t = Bound 0) t) fun reify'' (\<^term>\(+) :: real => _\ $ s $ t) = Add (reify' s, reify' t) | reify'' (\<^term>\(-) :: real => _\ $ s $ t) = Minus (reify' s, reify' t) | reify'' (\<^term>\(*) :: real => _\ $ s $ t) = Mult (reify' s, reify' t) | reify'' (\<^term>\(/) :: real => _\ $ s $ t) = Div (reify' s, reify' t) | reify'' (\<^term>\uminus :: real => _\ $ s) = Uminus (reify' s) | reify'' (\<^term>\inverse :: real => _\ $ s) = Inverse (reify' s) | reify'' (\<^term>\ln :: real => _\ $ (\<^term>\(powr) :: real => _\ $ s $ t)) = LnPowr (reify' s, reify' t) | reify'' (\<^term>\exp :: real => _\ $ (\<^term>\ln :: real => _\ $ s)) = ExpLn (reify' s) | reify'' (\<^term>\ln :: real => _\ $ s) = Ln (reify' s) | reify'' (\<^term>\exp :: real => _\ $ s) = Exp (reify' s) | reify'' (\<^term>\(powr) :: real => _\ $ s $ t) = (if is_const t then Powr' (reify' s, t) else Powr (reify' s, reify' t)) | reify'' (\<^term>\powr_nat :: real => _\ $ s $ t) = Powr_Nat (reify' s, reify' t) | reify'' (\<^term>\(^) :: real => _\ $ s $ t) = (if is_const' t then Power (reify' s, t) else raise TERM ("reify", [t'])) | reify'' (\<^term>\root\ $ s $ t) = (if is_const' s then Root (reify' t, s) else raise TERM ("reify", [t'])) | reify'' (\<^term>\abs :: real => _\ $ s) = Absolute (reify' s) | reify'' (\<^term>\sgn :: real => _\ $ s) = Sgn (reify' s) | reify'' (\<^term>\min :: real => _\ $ s $ t) = Min (reify' s, reify' t) | reify'' (\<^term>\max :: real => _\ $ s $ t) = Max (reify' s, reify' t) | reify'' (\<^term>\Multiseries_Expansion.rfloor\ $ s) = Floor (reify' s) | reify'' (\<^term>\Multiseries_Expansion.rceil\ $ s) = Ceiling (reify' s) | reify'' (\<^term>\Archimedean_Field.frac :: real \ real\ $ s) = Frac (reify' s) | reify'' (\<^term>\Multiseries_Expansion.rnatmod\ $ s $ t) = NatMod (reify' s, reify' t) | reify'' (\<^term>\sin :: real => _\ $ s) = Sin (reify' s) | reify'' (\<^term>\arctan :: real => _\ $ s) = ArcTan (reify' s) | reify'' (\<^term>\cos :: real => _\ $ s) = Cos (reify' s) | reify'' (Bound 0) = X | reify'' t = case reify_custom ctxt t of SOME ((name, t), ts) => let val args = map (reify_aux ctxt t') ts in Custom (name, t, args) end | NONE => raise TERM ("reify", [t']) and reify' t = if is_const t then ConstExpr t else reify'' t in case Envir.eta_long [] t of Abs (_, \<^typ>\real\, t'') => reify' t'' | _ => raise TERM ("reify", [t]) end fun reify ctxt t = let val thm = preproc_term_conv ctxt (Thm.cterm_of ctxt t) val rhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd in (reify_aux ctxt t rhs, thm) end fun reify_simple_aux ctxt t' t = let fun is_const t = fastype_of (Abs ("x", \<^typ>\real\, t)) = \<^typ>\real \ real\ andalso not (exists_subterm (fn t => t = Bound 0) t) fun is_const' t = not (exists_subterm (fn t => t = Bound 0) t) fun reify'' (\<^term>\(+) :: real => _\ $ s $ t) = Add (reify'' s, reify'' t) | reify'' (\<^term>\(-) :: real => _\ $ s $ t) = Minus (reify'' s, reify'' t) | reify'' (\<^term>\(*) :: real => _\ $ s $ t) = Mult (reify'' s, reify'' t) | reify'' (\<^term>\(/) :: real => _\ $ s $ t) = Div (reify'' s, reify'' t) | reify'' (\<^term>\uminus :: real => _\ $ s) = Uminus (reify'' s) | reify'' (\<^term>\inverse :: real => _\ $ s) = Inverse (reify'' s) | reify'' (\<^term>\ln :: real => _\ $ s) = Ln (reify'' s) | reify'' (\<^term>\exp :: real => _\ $ s) = Exp (reify'' s) | reify'' (\<^term>\(powr) :: real => _\ $ s $ t) = Powr (reify'' s, reify'' t) | reify'' (\<^term>\powr_nat :: real => _\ $ s $ t) = Powr_Nat (reify'' s, reify'' t) | reify'' (\<^term>\(^) :: real => _\ $ s $ t) = (if is_const' t then Power (reify'' s, t) else raise TERM ("reify", [t'])) | reify'' (\<^term>\root\ $ s $ t) = (if is_const' s then Root (reify'' t, s) else raise TERM ("reify", [t'])) | reify'' (\<^term>\abs :: real => _\ $ s) = Absolute (reify'' s) | reify'' (\<^term>\sgn :: real => _\ $ s) = Sgn (reify'' s) | reify'' (\<^term>\min :: real => _\ $ s $ t) = Min (reify'' s, reify'' t) | reify'' (\<^term>\max :: real => _\ $ s $ t) = Max (reify'' s, reify'' t) | reify'' (\<^term>\Multiseries_Expansion.rfloor\ $ s) = Floor (reify'' s) | reify'' (\<^term>\Multiseries_Expansion.rceil\ $ s) = Ceiling (reify'' s) | reify'' (\<^term>\Archimedean_Field.frac :: real \ real\ $ s) = Frac (reify'' s) | reify'' (\<^term>\Multiseries_Expansion.rnatmod\ $ s $ t) = NatMod (reify'' s, reify'' t) | reify'' (\<^term>\sin :: real => _\ $ s) = Sin (reify'' s) | reify'' (\<^term>\cos :: real => _\ $ s) = Cos (reify'' s) | reify'' (Bound 0) = X | reify'' t = if is_const t then ConstExpr t else case reify_custom ctxt t of SOME ((name, t), ts) => let val args = map (reify_aux ctxt t') ts in Custom (name, t, args) end | NONE => raise TERM ("reify", [t']) in case Envir.eta_long [] t of Abs (_, \<^typ>\real\, t'') => reify'' t'' | _ => raise TERM ("reify", [t]) end fun reify_simple ctxt t = let val thm = preproc_term_conv ctxt (Thm.cterm_of ctxt t) val rhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd in (reify_simple_aux ctxt t rhs, thm) end fun simple_print_const (Free (x, _)) = x | simple_print_const (\<^term>\uminus :: real => real\ $ a) = "(-" ^ simple_print_const a ^ ")" | simple_print_const (\<^term>\(+) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "+" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\(-) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "-" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\(*) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "*" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\inverse :: real => _\ $ a) = "(1 / " ^ simple_print_const a ^ ")" | simple_print_const (\<^term>\(/) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "/" ^ simple_print_const b ^ ")" | simple_print_const t = Int.toString (snd (HOLogic.dest_number t)) fun to_mathematica (Add (a, b)) = "(" ^ to_mathematica a ^ " + " ^ to_mathematica b ^ ")" | to_mathematica (Minus (a, b)) = "(" ^ to_mathematica a ^ " - " ^ to_mathematica b ^ ")" | to_mathematica (Mult (a, b)) = "(" ^ to_mathematica a ^ " * " ^ to_mathematica b ^ ")" | to_mathematica (Div (a, b)) = "(" ^ to_mathematica a ^ " / " ^ to_mathematica b ^ ")" | to_mathematica (Powr (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ ")" | to_mathematica (Powr_Nat (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ ")" | to_mathematica (Powr' (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica (ConstExpr b) ^ ")" | to_mathematica (LnPowr (a, b)) = "Log[" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ "]" | to_mathematica (ExpLn a) = "Exp[Ln[" ^ to_mathematica a ^ "]]" | to_mathematica (Power (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica (ConstExpr b) ^ ")" | to_mathematica (Root (a, \<^term>\2::real\)) = "Sqrt[" ^ to_mathematica a ^ "]" | to_mathematica (Root (a, b)) = "Surd[" ^ to_mathematica a ^ ", " ^ to_mathematica (ConstExpr b) ^ "]" | to_mathematica (Uminus a) = "(-" ^ to_mathematica a ^ ")" | to_mathematica (Inverse a) = "(1/(" ^ to_mathematica a ^ "))" | to_mathematica (Exp a) = "Exp[" ^ to_mathematica a ^ "]" | to_mathematica (Ln a) = "Log[" ^ to_mathematica a ^ "]" | to_mathematica (Sin a) = "Sin[" ^ to_mathematica a ^ "]" | to_mathematica (Cos a) = "Cos[" ^ to_mathematica a ^ "]" | to_mathematica (ArcTan a) = "ArcTan[" ^ to_mathematica a ^ "]" | to_mathematica (Absolute a) = "Abs[" ^ to_mathematica a ^ "]" | to_mathematica (Sgn a) = "Sign[" ^ to_mathematica a ^ "]" | to_mathematica (Min (a, b)) = "Min[" ^ to_mathematica a ^ ", " ^ to_mathematica b ^ "]" | to_mathematica (Max (a, b)) = "Max[" ^ to_mathematica a ^ ", " ^ to_mathematica b ^ "]" | to_mathematica (Floor a) = "Floor[" ^ to_mathematica a ^ "]" | to_mathematica (Ceiling a) = "Ceiling[" ^ to_mathematica a ^ "]" | to_mathematica (Frac a) = "Mod[" ^ to_mathematica a ^ ", 1]" | to_mathematica (ConstExpr t) = simple_print_const t | to_mathematica X = "X" (* TODO: correct translation of frac() for Maple and Sage *) fun to_maple (Add (a, b)) = "(" ^ to_maple a ^ " + " ^ to_maple b ^ ")" | to_maple (Minus (a, b)) = "(" ^ to_maple a ^ " - " ^ to_maple b ^ ")" | to_maple (Mult (a, b)) = "(" ^ to_maple a ^ " * " ^ to_maple b ^ ")" | to_maple (Div (a, b)) = "(" ^ to_maple a ^ " / " ^ to_maple b ^ ")" | to_maple (Powr (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (Powr_Nat (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (Powr' (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple (ConstExpr b) ^ ")" | to_maple (LnPowr (a, b)) = "ln(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (ExpLn a) = "ln(exp(" ^ to_maple a ^ "))" | to_maple (Power (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple (ConstExpr b) ^ ")" | to_maple (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_maple a ^ ")" | to_maple (Root (a, b)) = "root(" ^ to_maple a ^ ", " ^ to_maple (ConstExpr b) ^ ")" | to_maple (Uminus a) = "(-" ^ to_maple a ^ ")" | to_maple (Inverse a) = "(1/(" ^ to_maple a ^ "))" | to_maple (Exp a) = "exp(" ^ to_maple a ^ ")" | to_maple (Ln a) = "ln(" ^ to_maple a ^ ")" | to_maple (Sin a) = "sin(" ^ to_maple a ^ ")" | to_maple (Cos a) = "cos(" ^ to_maple a ^ ")" | to_maple (ArcTan a) = "arctan(" ^ to_maple a ^ ")" | to_maple (Absolute a) = "abs(" ^ to_maple a ^ ")" | to_maple (Sgn a) = "signum(" ^ to_maple a ^ ")" | to_maple (Min (a, b)) = "min(" ^ to_maple a ^ ", " ^ to_maple b ^ ")" | to_maple (Max (a, b)) = "max(" ^ to_maple a ^ ", " ^ to_maple b ^ ")" | to_maple (Floor a) = "floor(" ^ to_maple a ^ ")" | to_maple (Ceiling a) = "ceil(" ^ to_maple a ^ ")" | to_maple (Frac a) = "frac(" ^ to_maple a ^ ")" | to_maple (ConstExpr t) = simple_print_const t | to_maple X = "x" fun to_maxima (Add (a, b)) = "(" ^ to_maxima a ^ " + " ^ to_maxima b ^ ")" | to_maxima (Minus (a, b)) = "(" ^ to_maxima a ^ " - " ^ to_maxima b ^ ")" | to_maxima (Mult (a, b)) = "(" ^ to_maxima a ^ " * " ^ to_maxima b ^ ")" | to_maxima (Div (a, b)) = "(" ^ to_maxima a ^ " / " ^ to_maxima b ^ ")" | to_maxima (Powr (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Powr_Nat (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Powr' (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (ExpLn a) = "exp (log (" ^ to_maxima a ^ "))" | to_maxima (LnPowr (a, b)) = "log(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Power (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_maxima a ^ ")" | to_maxima (Root (a, b)) = to_maxima a ^ "^(1/" ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (Uminus a) = "(-" ^ to_maxima a ^ ")" | to_maxima (Inverse a) = "(1/(" ^ to_maxima a ^ "))" | to_maxima (Exp a) = "exp(" ^ to_maxima a ^ ")" | to_maxima (Ln a) = "log(" ^ to_maxima a ^ ")" | to_maxima (Sin a) = "sin(" ^ to_maxima a ^ ")" | to_maxima (Cos a) = "cos(" ^ to_maxima a ^ ")" | to_maxima (ArcTan a) = "atan(" ^ to_maxima a ^ ")" | to_maxima (Absolute a) = "abs(" ^ to_maxima a ^ ")" | to_maxima (Sgn a) = "signum(" ^ to_maxima a ^ ")" | to_maxima (Min (a, b)) = "min(" ^ to_maxima a ^ ", " ^ to_maxima b ^ ")" | to_maxima (Max (a, b)) = "max(" ^ to_maxima a ^ ", " ^ to_maxima b ^ ")" | to_maxima (Floor a) = "floor(" ^ to_maxima a ^ ")" | to_maxima (Ceiling a) = "ceil(" ^ to_maxima a ^ ")" | to_maxima (Frac a) = let val x = to_maxima a in "(" ^ x ^ " - floor(" ^ x ^ "))" end | to_maxima (ConstExpr t) = simple_print_const t | to_maxima X = "x" fun to_sympy (Add (a, b)) = "(" ^ to_sympy a ^ " + " ^ to_sympy b ^ ")" | to_sympy (Minus (a, b)) = "(" ^ to_sympy a ^ " - " ^ to_sympy b ^ ")" | to_sympy (Mult (a, b)) = "(" ^ to_sympy a ^ " * " ^ to_sympy b ^ ")" | to_sympy (Div (a, b)) = "(" ^ to_sympy a ^ " / " ^ to_sympy b ^ ")" | to_sympy (Powr (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Powr_Nat (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Powr' (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (ExpLn a) = "exp (log (" ^ to_sympy a ^ "))" | to_sympy (LnPowr (a, b)) = "log(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Power (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_sympy a ^ ")" | to_sympy (Root (a, b)) = "root(" ^ to_sympy a ^ ", " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (Uminus a) = "(-" ^ to_sympy a ^ ")" | to_sympy (Inverse a) = "(1/(" ^ to_sympy a ^ "))" | to_sympy (Exp a) = "exp(" ^ to_sympy a ^ ")" | to_sympy (Ln a) = "log(" ^ to_sympy a ^ ")" | to_sympy (Sin a) = "sin(" ^ to_sympy a ^ ")" | to_sympy (Cos a) = "cos(" ^ to_sympy a ^ ")" | to_sympy (ArcTan a) = "atan(" ^ to_sympy a ^ ")" | to_sympy (Absolute a) = "abs(" ^ to_sympy a ^ ")" | to_sympy (Sgn a) = "sign(" ^ to_sympy a ^ ")" | to_sympy (Min (a, b)) = "min(" ^ to_sympy a ^ ", " ^ to_sympy b ^ ")" | to_sympy (Max (a, b)) = "max(" ^ to_sympy a ^ ", " ^ to_sympy b ^ ")" | to_sympy (Floor a) = "floor(" ^ to_sympy a ^ ")" | to_sympy (Ceiling a) = "ceiling(" ^ to_sympy a ^ ")" | to_sympy (Frac a) = "frac(" ^ to_sympy a ^ ")" | to_sympy (ConstExpr t) = simple_print_const t | to_sympy X = "x" fun to_sage (Add (a, b)) = "(" ^ to_sage a ^ " + " ^ to_sage b ^ ")" | to_sage (Minus (a, b)) = "(" ^ to_sage a ^ " - " ^ to_sage b ^ ")" | to_sage (Mult (a, b)) = "(" ^ to_sage a ^ " * " ^ to_sage b ^ ")" | to_sage (Div (a, b)) = "(" ^ to_sage a ^ " / " ^ to_sage b ^ ")" | to_sage (Powr (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Powr_Nat (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Powr' (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage (ConstExpr b) ^ ")" | to_sage (ExpLn a) = "exp (log (" ^ to_sage a ^ "))" | to_sage (LnPowr (a, b)) = "log(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Power (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage (ConstExpr b) ^ ")" | to_sage (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_sage a ^ ")" | to_sage (Root (a, b)) = to_sage a ^ "^(1/" ^ to_sage (ConstExpr b) ^ ")" | to_sage (Uminus a) = "(-" ^ to_sage a ^ ")" | to_sage (Inverse a) = "(1/(" ^ to_sage a ^ "))" | to_sage (Exp a) = "exp(" ^ to_sage a ^ ")" | to_sage (Ln a) = "log(" ^ to_sage a ^ ")" | to_sage (Sin a) = "sin(" ^ to_sage a ^ ")" | to_sage (Cos a) = "cos(" ^ to_sage a ^ ")" | to_sage (ArcTan a) = "atan(" ^ to_sage a ^ ")" | to_sage (Absolute a) = "abs(" ^ to_sage a ^ ")" | to_sage (Sgn a) = "sign(" ^ to_sage a ^ ")" | to_sage (Min (a, b)) = "min(" ^ to_sage a ^ ", " ^ to_sage b ^ ")" | to_sage (Max (a, b)) = "max(" ^ to_sage a ^ ", " ^ to_sage b ^ ")" | to_sage (Floor a) = "floor(" ^ to_sage a ^ ")" | to_sage (Ceiling a) = "ceil(" ^ to_sage a ^ ")" | to_sage (Frac a) = "frac(" ^ to_sage a ^ ")" | to_sage (ConstExpr t) = simple_print_const t | to_sage X = "x" fun reify_mathematica ctxt = to_mathematica o fst o reify_simple ctxt fun reify_maple ctxt = to_maple o fst o reify_simple ctxt fun reify_maxima ctxt = to_maxima o fst o reify_simple ctxt fun reify_sympy ctxt = to_sympy o fst o reify_simple ctxt fun reify_sage ctxt = to_sage o fst o reify_simple ctxt fun limit_mathematica s = "Limit[" ^ s ^ ", X -> Infinity]" fun limit_maple s = "limit(" ^ s ^ ", x = infinity);" fun limit_maxima s = "limit(" ^ s ^ ", x, inf);" fun limit_sympy s = "limit(" ^ s ^ ", x, oo)" fun limit_sage s = "limit(" ^ s ^ ", x = Infinity)" end diff --git a/src/HOL/Statespace/state_space.ML b/src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML +++ b/src/HOL/Statespace/state_space.ML @@ -1,734 +1,736 @@ (* Title: HOL/Statespace/state_space.ML Author: Norbert Schirmer, TU Muenchen, 2007 Author: Norbert Schirmer, Apple, 2021 *) signature STATE_SPACE = sig val distinct_compsN : string val getN : string val putN : string val injectN : string val namespaceN : string val projectN : string val valuetypesN : string val namespace_definition : bstring -> typ -> (xstring, string) Expression.expr * (binding * string option * mixfix) list -> string list -> string list -> theory -> theory val define_statespace : string list -> string -> ((string * bool) * (string list * bstring * (string * string) list)) list -> (string * string) list -> theory -> theory val define_statespace_i : string option -> string list -> string -> ((string * bool) * (typ list * bstring * (string * string) list)) list -> (string * typ) list -> theory -> theory val statespace_decl : ((string list * bstring) * (((string * bool) * (string list * xstring * (bstring * bstring) list)) list * (bstring * string) list)) parser val neq_x_y : Proof.context -> term -> term -> thm option val distinctNameSolver : Simplifier.solver val distinctTree_tac : Proof.context -> int -> tactic val distinct_simproc : Simplifier.simproc val is_statespace : Context.generic -> xstring -> bool val get_comp' : Context.generic -> string -> (typ * string list) option val get_comp : Context.generic -> string -> (typ * string) option (* legacy wrapper, eventually replace by primed version *) val get_comps : Context.generic -> (typ * string list) Termtab.table val silent : bool Config.T val gen_lookup_tr : Proof.context -> term -> string -> term val lookup_swap_tr : Proof.context -> term list -> term val lookup_tr : Proof.context -> term list -> term val lookup_tr' : Proof.context -> term list -> term val gen'_update_tr : bool -> bool -> Proof.context -> string -> term -> term -> term val gen_update_tr : (* legacy wrapper, eventually replace by primed version *) bool -> Proof.context -> string -> term -> term -> term val update_tr : Proof.context -> term list -> term val update_tr' : Proof.context -> term list -> term val trace_data: Context.generic -> unit end; structure StateSpace : STATE_SPACE = struct (* Names *) val distinct_compsN = "distinct_names" val namespaceN = "_namespace" val valuetypesN = "_valuetypes" val projectN = "project" val injectN = "inject" val getN = "get" val putN = "put" val project_injectL = "StateSpaceLocale.project_inject"; (* Library *) fun fold1 f xs = fold f (tl xs) (hd xs) fun fold1' f [] x = x | fold1' f xs _ = fold1 f xs fun sorted_subset eq [] ys = true | sorted_subset eq (x::xs) [] = false | sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys else sorted_subset eq (x::xs) ys; fun comps_of_distinct_thm thm = Thm.prop_of thm |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free) |> sort_strings; fun insert_tagged_distinct_thms tagged_thm tagged_thms = let fun ins t1 [] = [t1] | ins (t1 as (names1, _)) ((t2 as (names2, _))::thms) = if Ord_List.subset string_ord (names1, names2) then t2::thms else if Ord_List.subset string_ord (names2, names1) then ins t1 thms else t2 :: ins t1 thms in ins tagged_thm tagged_thms end fun join_tagged_distinct_thms tagged_thms1 tagged_thms2 = tagged_thms1 |> fold insert_tagged_distinct_thms tagged_thms2 fun tag_distinct_thm thm = (comps_of_distinct_thm thm, thm) val tag_distinct_thms = map tag_distinct_thm fun join_distinct_thms (thms1, thms2) = if pointer_eq (thms1, thms2) then thms1 else join_tagged_distinct_thms (tag_distinct_thms thms1) (tag_distinct_thms thms2) |> map snd fun insert_distinct_thm thm thms = join_distinct_thms (thms, [thm]) fun join_declinfo_entry name (T1:typ, names1:string list) (T2, names2) = let fun typ_info T names = @{make_string} T ^ " " ^ Pretty.string_of (Pretty.str_list "(" ")" names); in if T1 = T2 then (T1, distinct (op =) (names1 @ names2)) else error ("statespace component '" ^ name ^ "' disagrees on type:\n " ^ typ_info T1 names1 ^ " vs. " ^ typ_info T2 names2 ) end fun guess_name (Free (x,_)) = x | guess_name _ = "unknown" val join_declinfo = Termtab.join (fn t => uncurry (join_declinfo_entry (guess_name t))) (* A component might appear in *different* statespaces within the same context. However, it must be mapped to the same type. Note that this information is currently only properly maintained within contexts where all components are actually "fixes" and not arbitrary terms. Moreover, on the theory level the info stays empty. This means that on the theory level we do not make use of the syntax and the tree-based distinct simprocs. N.B: It might still make sense (from a performance perspective) to overcome this limitation and even use the simprocs when a statespace is interpreted for concrete names like HOL-strings. Once the distinct-theorem is proven by the interpretation the simproc can use the positions in the tree to derive distinctness of two strings vs. HOL-string comparison. *) type statespace_info = {args: (string * sort) list, (* type arguments *) parents: (typ list * string * string option list) list, (* type instantiation, state-space name, component renamings *) components: (string * typ) list, types: typ list (* range types of state space *) }; structure Data = Generic_Data ( type T = (typ * string list) Termtab.table * (*declinfo: type, names of statespaces of component*) thm list Symtab.table * (*distinctthm: minimal list of maximal distinctness-assumptions for a component name*) statespace_info Symtab.table; val empty = (Termtab.empty, Symtab.empty, Symtab.empty); fun merge ((declinfo1, distinctthm1, statespaces1), (declinfo2, distinctthm2, statespaces2)) = (join_declinfo (declinfo1, declinfo2), Symtab.join (K join_distinct_thms) (distinctthm1, distinctthm2), Symtab.merge (K true) (statespaces1, statespaces2)); ); val get_declinfo = #1 o Data.get val get_distinctthm = #2 o Data.get val get_statespaces = #3 o Data.get val map_declinfo = Data.map o @{apply 3(1)} val map_distinctthm = Data.map o @{apply 3(2)} val map_statespaces = Data.map o @{apply 3(3)} fun trace_data context = tracing ("StateSpace.Data: " ^ @{make_string} {declinfo = get_declinfo context, distinctthm = get_distinctthm context, statespaces = get_statespaces context}) fun update_declinfo (n,v) = map_declinfo (fn declinfo => let val vs = apsnd single v in Termtab.map_default (n, vs) (join_declinfo_entry (guess_name n) vs) declinfo end); fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun prove_interpretation_in ctxt_tac (name, expr) thy = thy |> Interpretation.global_sublocale_cmd (name, Position.none) (expression_no_pos expr) [] |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), Position.no_range), NONE) |> Proof_Context.theory_of fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) Binding.empty [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun is_statespace context name = Symtab.defined (get_statespaces context) (Locale.intern (Context.theory_of context) name) fun add_statespace name args parents components types = map_statespaces (Symtab.update_new (name, {args=args,parents=parents, components=components,types=types})) val get_statespace = Symtab.lookup o get_statespaces val the_statespace = the oo get_statespace fun mk_free ctxt name = if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name then let val n' = Variable.intern_fixed ctxt name |> perhaps Long_Name.dest_hidden; val free = Free (n', Proof_Context.infer_type ctxt (n', dummyT)) in SOME (free) end else (tracing ("variable not fixed or declared: " ^ name); NONE) fun get_dist_thm context name = Symtab.lookup_list (get_distinctthm context) name |> map (Thm.transfer'' context) fun get_dist_thm2 ctxt x y = (let val dist_thms = [x, y] |> map (#1 o dest_Free) |> maps (get_dist_thm (Context.Proof ctxt)); fun get_paths dist_thm = let val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val tree = Thm.term_of ctree; val x_path = the (DistinctTreeProver.find_tree x tree); val y_path = the (DistinctTreeProver.find_tree y tree); in SOME (dist_thm, x_path, y_path) end handle Option.Option => NONE val result = get_first get_paths dist_thms in result end) fun get_comp' context name = mk_free (Context.proof_of context) name |> Option.mapPartial (fn t => let val declinfo = get_declinfo context in case Termtab.lookup declinfo t of NONE => (* during syntax phase, types of fixes might not yet be constrained completely *) AList.lookup (fn (x, Free (n,_)) => n = x | _ => false) (Termtab.dest declinfo) name | some => some end) (* legacy wrapper *) fun get_comp ctxt name = get_comp' ctxt name |> Option.map (apsnd (fn ns => if null ns then "" else hd ns)) val get_comps = get_declinfo (*** Tactics ***) fun neq_x_y ctxt x y = (let val (dist_thm, x_path, y_path) = the (get_dist_thm2 ctxt x y); val thm = DistinctTreeProver.distinctTreeProver ctxt dist_thm x_path y_path; in SOME thm end handle Option.Option => NONE) fun distinctTree_tac ctxt = SUBGOAL (fn (goal, i) => (case goal of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ (x as Free _) $ (y as Free _))) => (case neq_x_y ctxt x y of SOME neq => resolve_tac ctxt [neq] i | NONE => no_tac) | _ => no_tac)); val distinctNameSolver = mk_solver "distinctNameSolver" distinctTree_tac; val distinct_simproc = Simplifier.make_simproc \<^context> "StateSpace.distinct_simproc" {lhss = [\<^term>\x = y\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\,_) $ (x as Free _) $ (y as Free _) => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) (neq_x_y ctxt x y) | _ => NONE)}; fun interprete_parent name dist_thm_name parent_expr thy = let fun solve_tac ctxt = CSUBGOAL (fn (goal, i) => let val distinct_thm = Proof_Context.get_thm ctxt dist_thm_name; val rule = DistinctTreeProver.distinct_implProver ctxt distinct_thm goal; in resolve_tac ctxt [rule] i end); fun tac ctxt = Locale.intro_locales_tac {strict = true, eager = true} ctxt [] THEN ALLGOALS (solve_tac ctxt); in thy |> prove_interpretation_in tac (name, parent_expr) end; fun namespace_definition name nameT parent_expr parent_comps new_comps thy = let val all_comps = parent_comps @ new_comps; val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps); val dist_thm_name = distinct_compsN; val dist_thm_full_name = dist_thm_name; fun type_attr phi = Thm.declaration_attribute (fn thm => fn context => (case context of Context.Theory _ => context | Context.Proof ctxt => let val declinfo = get_declinfo context val tt = get_distinctthm context; val all_names = comps_of_distinct_thm thm; val thm0 = Thm.trim_context thm; fun upd name = Symtab.map_default (name, [thm0]) (insert_distinct_thm thm0) val tt' = tt |> fold upd all_names; val context' = Context_Position.set_visible false ctxt addsimprocs [distinct_simproc] |> Context_Position.restore_visible ctxt |> Context.Proof |> map_declinfo (K declinfo) |> map_distinctthm (K tt'); in context' end)); - val attr = Attrib.internal type_attr; + val attr = Attrib.internal \<^here> type_attr; val assume = ((Binding.name dist_thm_name, [attr]), [(HOLogic.Trueprop $ (Const (\<^const_name>\all_distinct\, Type (\<^type_name>\tree\, [nameT]) --> HOLogic.boolT) $ DistinctTreeProver.mk_tree (fn n => Free (n, nameT)) nameT (sort fast_string_ord all_comps)), [])]); in thy |> add_locale name ([], vars) [Element.Assumes [assume]] |> Proof_Context.theory_of |> interprete_parent name dist_thm_full_name parent_expr end; fun encode_dot x = if x = #"." then #"_" else x; fun encode_type (TFree (s, _)) = s | encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i | encode_type (Type (n,Ts)) = let val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) ""; val n' = String.map encode_dot n; in if Ts'="" then n' else Ts' ^ "_" ^ n' end; fun project_name T = projectN ^"_"^encode_type T; fun inject_name T = injectN ^"_"^encode_type T; fun add_declaration name decl thy = thy |> Named_Target.init [] name - |> (fn lthy => Local_Theory.declaration {syntax = true, pervasive = false} (decl lthy) lthy) + |> (fn lthy => + Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} + (decl lthy) lthy) |> Local_Theory.exit_global; fun parent_components thy (Ts, pname, renaming) = let fun rename [] xs = xs | rename (NONE::rs) (x::xs) = x::rename rs xs | rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs; val {args, parents, components, ...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val parent_comps = maps (fn (Ts',n,rs) => parent_components thy (map subst Ts', n, rs)) parents; val all_comps = rename renaming (parent_comps @ map (apsnd subst) components); in all_comps end; fun statespace_definition state_type args name parents parent_comps components thy = let val full_name = Sign.full_bname thy name; val all_comps = parent_comps @ components; val components' = map (fn (n,T) => (n,(T,full_name))) components; fun parent_expr (prefix, (_, n, rs)) = (suffix namespaceN n, (prefix, (Expression.Positional rs,[]))); val parents_expr = map parent_expr parents; fun distinct_types Ts = let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty; in map fst (Typtab.dest tab) end; val Ts = distinct_types (map snd all_comps); val arg_names = map fst args; val valueN = singleton (Name.variant_list arg_names) "'value"; val nameN = singleton (Name.variant_list (valueN :: arg_names)) "'name"; val valueT = TFree (valueN, Sign.defaultS thy); val nameT = TFree (nameN, Sign.defaultS thy); val stateT = nameT --> valueT; fun projectT T = valueT --> T; fun injectT T = T --> valueT; val locinsts = map (fn T => (project_injectL, ((encode_type T,false),(Expression.Positional [SOME (Free (project_name T,projectT T)), SOME (Free ((inject_name T,injectT T)))],[])))) Ts; val locs = maps (fn T => [(Binding.name (project_name T),NONE,NoSyn), (Binding.name (inject_name T),NONE,NoSyn)]) Ts; val constrains = maps (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts; fun interprete_parent_valuetypes (prefix, (Ts, pname, _)) thy = let val {args,types,...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val pars = maps ((fn T => [project_name T,inject_name T]) o subst) types; val expr = ([(suffix valuetypesN name, (prefix, (Expression.Positional (map SOME pars),[])))],[]); in prove_interpretation_in (fn ctxt => ALLGOALS (solve_tac ctxt (Assumption.all_prems_of ctxt))) (suffix valuetypesN name, expr) thy end; fun interprete_parent (prefix, (_, pname, rs)) = let val expr = ([(pname, (prefix, (Expression.Positional rs,[])))],[]) in prove_interpretation_in (fn ctxt => Locale.intro_locales_tac {strict = true, eager = false} ctxt []) (full_name, expr) end; fun declare_declinfo updates lthy phi ctxt = let fun upd_prf ctxt = let fun upd (n,v) = let val nT = Proof_Context.infer_type (Local_Theory.target_of lthy) (n, dummyT) in Context.proof_map (update_declinfo (Morphism.term phi (Free (n,nT)),v)) end; val ctxt' = ctxt |> fold upd updates in ctxt' end; in Context.mapping I upd_prf ctxt end; fun string_of_typ T = Print_Mode.setmp [] (Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy))) T; val fixestate = (case state_type of NONE => [] | SOME s => let val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)]; val cs = Element.Constrains (map (fn (n,T) => (n,string_of_typ T)) ((map (fn (n,_) => (n,nameT)) all_comps) @ constrains)) in [fx,cs] end ) in thy |> namespace_definition (suffix namespaceN name) nameT (parents_expr,[]) (map fst parent_comps) (map fst components) |> Context.theory_map (add_statespace full_name args (map snd parents) components []) |> add_locale (suffix valuetypesN name) (locinsts,locs) [] |> Proof_Context.theory_of |> fold interprete_parent_valuetypes parents |> add_locale_cmd name ([(suffix namespaceN full_name ,(("",false),(Expression.Named [],[]))), (suffix valuetypesN full_name,(("",false),(Expression.Named [],[])))],[]) fixestate |> Proof_Context.theory_of |> fold interprete_parent parents |> add_declaration full_name (declare_declinfo components') end; (* prepare arguments *) fun read_typ ctxt raw_T env = let val ctxt' = fold (Variable.declare_typ o TFree) env ctxt; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun cert_typ ctxt raw_T env = let val thy = Proof_Context.theory_of ctxt; val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg; val env' = Term.add_tfreesT T env; in (T, env') end; fun gen_define_statespace prep_typ state_space args name parents comps thy = let (* - args distinct - only args may occur in comps and parent-instantiations - number of insts must match parent args - no duplicate renamings - renaming should occur in namespace *) val _ = writeln ("Defining statespace " ^ quote name ^ " ..."); val ctxt = Proof_Context.init_global thy; fun add_parent (prefix, (Ts, pname, rs)) env = let val prefix' = (case prefix of ("", mandatory) => (pname, mandatory) | _ => prefix); val full_pname = Sign.full_bname thy pname; val {args,components,...} = (case get_statespace (Context.Theory thy) full_pname of SOME r => r | NONE => error ("Undefined statespace " ^ quote pname)); val (Ts',env') = fold_map (prep_typ ctxt) Ts env handle ERROR msg => cat_error msg ("The error(s) above occurred in parent statespace specification " ^ quote pname); val err_insts = if length args <> length Ts' then ["number of type instantiation(s) does not match arguments of parent statespace " ^ quote pname] else []; val rnames = map fst rs val err_dup_renamings = (case duplicates (op =) rnames of [] => [] | dups => ["Duplicate renaming(s) for " ^ commas dups]) val cnames = map fst components; val err_rename_unknowns = (case subtract (op =) cnames rnames of [] => [] | rs => ["Unknown components " ^ commas rs]); val rs' = map (AList.lookup (op =) rs o fst) components; val errs =err_insts @ err_dup_renamings @ err_rename_unknowns in if null errs then ((prefix', (Ts', full_pname, rs')), env') else error (cat_lines (errs @ ["in parent statespace " ^ quote pname])) end; val (parents',env) = fold_map add_parent parents []; val err_dup_args = (case duplicates (op =) args of [] => [] | dups => ["Duplicate type argument(s) " ^ commas dups]); val err_dup_components = (case duplicates (op =) (map fst comps) of [] => [] | dups => ["Duplicate state-space components " ^ commas dups]); fun prep_comp (n,T) env = let val (T', env') = prep_typ ctxt T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n,T'), env') end; val (comps',env') = fold_map prep_comp comps env; val err_extra_frees = (case subtract (op =) args (map fst env') of [] => [] | extras => ["Extra free type variable(s) " ^ commas extras]); val defaultS = Sign.defaultS thy; val args' = map (fn x => (x, AList.lookup (op =) env x |> the_default defaultS)) args; fun fst_eq ((x:string,_),(y,_)) = x = y; fun snd_eq ((_,t:typ),(_,u)) = t = u; val raw_parent_comps = maps (parent_components thy o snd) parents'; fun check_type (n,T) = (case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of [] => [] | [_] => [] | rs => ["Different types for component " ^ quote n ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o snd) rs)]) val err_dup_types = maps check_type (duplicates fst_eq raw_parent_comps) val parent_comps = distinct (fst_eq) raw_parent_comps; val all_comps = parent_comps @ comps'; val err_comp_in_parent = (case duplicates (op =) (map fst all_comps) of [] => [] | xs => ["Components already defined in parents: " ^ commas_quote xs]); val errs = err_dup_args @ err_dup_components @ err_extra_frees @ err_dup_types @ err_comp_in_parent; in if null errs then thy |> statespace_definition state_space args' name parents' parent_comps comps' else error (cat_lines errs) end handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name); val define_statespace = gen_define_statespace read_typ NONE; val define_statespace_i = gen_define_statespace cert_typ; (*** parse/print - translations ***) val silent = Attrib.setup_config_bool \<^binding>\statespace_silent\ (K false); fun gen_lookup_tr ctxt s n = (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.free (project_name T) $ Syntax.free n $ s | NONE => if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ s else raise TERM ("StateSpace.gen_lookup_tr: component " ^ quote n ^ " not defined", [])); fun lookup_tr ctxt [s, x] = (case Term_Position.strip_positions x of Free (n,_) => gen_lookup_tr ctxt s n | _ => raise Match); fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n; fun lookup_tr' ctxt [_ $ Free (prj, _), n as (_ $ Free (name, _)), s] = (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if prj = project_name T then Syntax.const "_statespace_lookup" $ s $ n else raise Match | NONE => raise Match) | lookup_tr' _ _ = raise Match; fun gen'_update_tr const_val id ctxt n v s = let fun pname T = if id then \<^const_name>\Fun.id\ else project_name T; fun iname T = if id then \<^const_name>\Fun.id\ else inject_name T; val v = if const_val then (Syntax.const \<^const_name>\K_statefun\ $ v) else v in (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.update\ $ Syntax.free (pname T) $ Syntax.free (iname T) $ Syntax.free n $ v $ s | NONE => if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.update\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ v $ s else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined", [])) end; val gen_update_tr = gen'_update_tr true fun update_tr ctxt [s, x, v] = (case Term_Position.strip_positions x of Free (n, _) => gen'_update_tr true false ctxt n v s | _ => raise Match); fun update_tr' ctxt [_ $ Free (prj, _), _ $ Free (inj, _), n as (_ $ Free (name, _)), (Const (k, _) $ v), s] = if Long_Name.base_name k = Long_Name.base_name \<^const_name>\K_statefun\ then (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if inj = inject_name T andalso prj = project_name T then Syntax.const "_statespace_update" $ s $ n $ v else raise Match | NONE => raise Match) else raise Match | update_tr' _ _ = raise Match; (*** outer syntax *) local val type_insts = Parse.typ >> single || \<^keyword>\(\ |-- Parse.!!! (Parse.list1 Parse.typ --| \<^keyword>\)\) val comp = Parse.name -- (\<^keyword>\::\ |-- Parse.!!! Parse.typ); fun plus1_unless test scan = scan ::: Scan.repeat (\<^keyword>\+\ |-- Scan.unless test (Parse.!!! scan)); val mapsto = \<^keyword>\=\; val rename = Parse.name -- (mapsto |-- Parse.name); val renames = Scan.optional (\<^keyword>\[\ |-- Parse.!!! (Parse.list1 rename --| \<^keyword>\]\)) []; val parent = Parse_Spec.locale_prefix -- ((type_insts -- Parse.name) || (Parse.name >> pair [])) -- renames >> (fn ((prefix, (insts, name)), renames) => (prefix, (insts, name, renames))); in val statespace_decl = Parse.type_args -- Parse.name -- (\<^keyword>\=\ |-- ((Scan.repeat1 comp >> pair []) || (plus1_unless comp parent -- Scan.optional (\<^keyword>\+\ |-- Parse.!!! (Scan.repeat1 comp)) []))); val _ = Outer_Syntax.command \<^command_keyword>\statespace\ "define state-space as locale context" (statespace_decl >> (fn ((args, name), (parents, comps)) => Toplevel.theory (define_statespace args name parents comps))); end; end; diff --git a/src/HOL/Tools/BNF/bnf_def.ML b/src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML +++ b/src/HOL/Tools/BNF/bnf_def.ML @@ -1,2187 +1,2187 @@ (* Title: HOL/Tools/BNF/bnf_def.ML Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Author: Jan van Brügge, TU Muenchen Copyright 2012, 2013, 2014, 2022 Definition of bounded natural functors. *) signature BNF_DEF = sig type bnf type nonemptiness_witness = {I: int list, wit: term, prop: thm list} val morph_bnf: morphism -> bnf -> bnf val morph_bnf_defs: morphism -> bnf -> bnf val permute_deads: (typ list -> typ list) -> bnf -> bnf val transfer_bnf: theory -> bnf -> bnf val bnf_of: Proof.context -> string -> bnf option val bnf_of_global: theory -> string -> bnf option val bnf_interpretation: string -> (bnf -> local_theory -> local_theory) -> theory -> theory val interpret_bnf: (string -> bool) -> bnf -> local_theory -> local_theory val register_bnf_raw: string -> bnf -> local_theory -> local_theory val register_bnf: (string -> bool) -> string -> bnf -> local_theory -> local_theory val name_of_bnf: bnf -> binding val T_of_bnf: bnf -> typ val live_of_bnf: bnf -> int val lives_of_bnf: bnf -> typ list val dead_of_bnf: bnf -> int val deads_of_bnf: bnf -> typ list val bd_of_bnf: bnf -> term val nwits_of_bnf: bnf -> int val mapN: string val predN: string val relN: string val setN: string val mk_setN: int -> string val mk_witN: int -> string val map_of_bnf: bnf -> term val pred_of_bnf: bnf -> term val rel_of_bnf: bnf -> term val sets_of_bnf: bnf -> term list val mk_T_of_bnf: typ list -> typ list -> bnf -> typ val mk_bd_of_bnf: typ list -> typ list -> bnf -> term val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_pred_of_bnf: typ list -> typ list -> bnf -> term val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list val bd_Card_order_of_bnf: bnf -> thm val bd_Cinfinite_of_bnf: bnf -> thm val bd_Cnotzero_of_bnf: bnf -> thm val bd_card_order_of_bnf: bnf -> thm val bd_cinfinite_of_bnf: bnf -> thm val bd_regularCard_of_bnf: bnf -> thm val collect_set_map_of_bnf: bnf -> thm val in_bd_of_bnf: bnf -> thm val in_cong_of_bnf: bnf -> thm val in_mono_of_bnf: bnf -> thm val in_rel_of_bnf: bnf -> thm val inj_map_of_bnf: bnf -> thm val inj_map_strong_of_bnf: bnf -> thm val le_rel_OO_of_bnf: bnf -> thm val map_comp0_of_bnf: bnf -> thm val map_comp_of_bnf: bnf -> thm val map_cong0_of_bnf: bnf -> thm val map_cong_of_bnf: bnf -> thm val map_cong_pred_of_bnf: bnf -> thm val map_cong_simp_of_bnf: bnf -> thm val map_def_of_bnf: bnf -> thm val map_id0_of_bnf: bnf -> thm val map_id_of_bnf: bnf -> thm val map_ident0_of_bnf: bnf -> thm val map_ident_of_bnf: bnf -> thm val map_transfer_of_bnf: bnf -> thm val pred_cong0_of_bnf: bnf -> thm val pred_cong_of_bnf: bnf -> thm val pred_cong_simp_of_bnf: bnf -> thm val pred_def_of_bnf: bnf -> thm val pred_map_of_bnf: bnf -> thm val pred_mono_strong0_of_bnf: bnf -> thm val pred_mono_strong_of_bnf: bnf -> thm val pred_mono_of_bnf: bnf -> thm val pred_set_of_bnf: bnf -> thm val pred_rel_of_bnf: bnf -> thm val pred_transfer_of_bnf: bnf -> thm val pred_True_of_bnf: bnf -> thm val rel_Grp_of_bnf: bnf -> thm val rel_OO_Grp_of_bnf: bnf -> thm val rel_OO_of_bnf: bnf -> thm val rel_cong0_of_bnf: bnf -> thm val rel_cong_of_bnf: bnf -> thm val rel_cong_simp_of_bnf: bnf -> thm val rel_conversep_of_bnf: bnf -> thm val rel_def_of_bnf: bnf -> thm val rel_eq_of_bnf: bnf -> thm val rel_flip_of_bnf: bnf -> thm val rel_map_of_bnf: bnf -> thm list val rel_mono_of_bnf: bnf -> thm val rel_mono_strong0_of_bnf: bnf -> thm val rel_mono_strong_of_bnf: bnf -> thm val rel_eq_onp_of_bnf: bnf -> thm val rel_refl_of_bnf: bnf -> thm val rel_refl_strong_of_bnf: bnf -> thm val rel_reflp_of_bnf: bnf -> thm val rel_symp_of_bnf: bnf -> thm val rel_transfer_of_bnf: bnf -> thm val rel_transp_of_bnf: bnf -> thm val set_bd_of_bnf: bnf -> thm list val set_defs_of_bnf: bnf -> thm list val set_map0_of_bnf: bnf -> thm list val set_map_of_bnf: bnf -> thm list val set_transfer_of_bnf: bnf -> thm list val wit_thms_of_bnf: bnf -> thm list val wit_thmss_of_bnf: bnf -> thm list list val mk_map: int -> typ list -> typ list -> term -> term val mk_pred: typ list -> term -> term val mk_rel: int -> typ list -> typ list -> term -> term val mk_set: typ list -> term -> term val build_map: Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_set: Proof.context -> typ -> typ -> term val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list -> 'a list val mk_witness: int list * term -> thm list -> nonemptiness_witness val mk_wit_goals: term list -> term list -> term list -> int list * term -> term list val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list val wits_of_bnf: bnf -> nonemptiness_witness list val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline datatype fact_policy = Dont_Note | Note_Some | Note_All val bnf_internals: bool Config.T val bnf_timing: bool Config.T val user_policy: fact_policy -> Proof.context -> fact_policy val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> local_theory -> bnf * local_theory val note_bnf_defs: bnf -> local_theory -> bnf * local_theory val print_bnfs: Proof.context -> unit val prepare_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option) * 'b option -> Proof.context -> string * term list * ((Proof.context -> thm list -> tactic) option * term list list) * ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) * local_theory * thm list val define_bnf_consts: inline_policy -> fact_policy -> bool -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> ((typ list * typ list * typ list * typ) * (term * term list * term * (int list * term) list * term * term) * (thm * thm list * thm * thm list * thm * thm) * ((typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term -> term) * (typ list -> typ list -> typ -> typ) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term))) * local_theory val bnf_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> tactic) list -> (Proof.context -> tactic) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> bnf * local_theory val bnf_cmd: (((((((binding * string) * string) * string list) * string) * string list) * string option) * string option) * (Proof.context -> Plugin_Name.filter) -> Proof.context -> Proof.state end; structure BNF_Def : BNF_DEF = struct open BNF_Util open BNF_Tactics open BNF_Def_Tactics val fundefcong_attrs = @{attributes [fundef_cong]}; val mono_attrs = @{attributes [mono]}; type axioms = { map_id0: thm, map_comp0: thm, map_cong0: thm, set_map0: thm list, bd_card_order: thm, bd_cinfinite: thm, bd_regularCard: thm, set_bd: thm list, le_rel_OO: thm, rel_OO_Grp: thm, pred_set: thm }; fun mk_axioms' ((((((((((id, comp), cong), map), c_o), cinf), creg), set_bd), le_rel_OO), rel), pred) = {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o, bd_cinfinite = cinf, bd_regularCard = creg, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel, pred_set = pred}; fun dest_cons [] = raise List.Empty | dest_cons (x :: xs) = (x, xs); fun mk_axioms n thms = thms |> map the_single |> dest_cons ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||> the_single |> mk_axioms'; fun zip_axioms mid mcomp mcong smap bdco bdinf bdreg sbd le_rel_OO rel pred = [mid, mcomp, mcong] @ smap @ [bdco, bdinf, bdreg] @ sbd @ [le_rel_OO, rel, pred]; fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, bd_regularCard, set_bd, le_rel_OO, rel_OO_Grp, pred_set} = {map_id0 = f map_id0, map_comp0 = f map_comp0, map_cong0 = f map_cong0, set_map0 = map f set_map0, bd_card_order = f bd_card_order, bd_cinfinite = f bd_cinfinite, bd_regularCard = f bd_regularCard, set_bd = map f set_bd, le_rel_OO = f le_rel_OO, rel_OO_Grp = f rel_OO_Grp, pred_set = f pred_set}; val morph_axioms = map_axioms o Morphism.thm; type defs = { map_def: thm, set_defs: thm list, rel_def: thm, pred_def: thm } fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred}; fun map_defs f {map_def, set_defs, rel_def, pred_def} = {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, pred_def = f pred_def}; val morph_defs = map_defs o Morphism.thm; type facts = { bd_Card_order: thm, bd_Cinfinite: thm, bd_Cnotzero: thm, collect_set_map: thm lazy, in_bd: thm lazy, in_cong: thm lazy, in_mono: thm lazy, in_rel: thm lazy, inj_map: thm lazy, inj_map_strong: thm lazy, map_comp: thm lazy, map_cong: thm lazy, map_cong_simp: thm lazy, map_cong_pred: thm lazy, map_id: thm lazy, map_ident0: thm lazy, map_ident: thm lazy, map_ident_strong: thm lazy, map_transfer: thm lazy, rel_eq: thm lazy, rel_flip: thm lazy, set_map: thm lazy list, rel_cong0: thm lazy, rel_cong: thm lazy, rel_cong_simp: thm lazy, rel_map: thm list lazy, rel_mono: thm lazy, rel_mono_strong0: thm lazy, rel_mono_strong: thm lazy, set_transfer: thm list lazy, rel_Grp: thm lazy, rel_conversep: thm lazy, rel_OO: thm lazy, rel_refl: thm lazy, rel_refl_strong: thm lazy, rel_reflp: thm lazy, rel_symp: thm lazy, rel_transp: thm lazy, rel_transfer: thm lazy, rel_eq_onp: thm lazy, pred_transfer: thm lazy, pred_True: thm lazy, pred_map: thm lazy, pred_rel: thm lazy, pred_mono_strong0: thm lazy, pred_mono_strong: thm lazy, pred_mono: thm lazy, pred_cong0: thm lazy, pred_cong: thm lazy, pred_cong_simp: thm lazy }; fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp = { bd_Card_order = bd_Card_order, bd_Cinfinite = bd_Cinfinite, bd_Cnotzero = bd_Cnotzero, collect_set_map = collect_set_map, in_bd = in_bd, in_cong = in_cong, in_mono = in_mono, in_rel = in_rel, inj_map = inj_map, inj_map_strong = inj_map_strong, map_comp = map_comp, map_cong = map_cong, map_cong_simp = map_cong_simp, map_cong_pred = map_cong_pred, map_id = map_id, map_ident0 = map_ident0, map_ident = map_ident, map_ident_strong = map_ident_strong, map_transfer = map_transfer, rel_eq = rel_eq, rel_flip = rel_flip, set_map = set_map, rel_cong0 = rel_cong0, rel_cong = rel_cong, rel_cong_simp = rel_cong_simp, rel_map = rel_map, rel_mono = rel_mono, rel_mono_strong0 = rel_mono_strong0, rel_mono_strong = rel_mono_strong, rel_transfer = rel_transfer, rel_Grp = rel_Grp, rel_conversep = rel_conversep, rel_OO = rel_OO, rel_refl = rel_refl, rel_refl_strong = rel_refl_strong, rel_reflp = rel_reflp, rel_symp = rel_symp, rel_transp = rel_transp, set_transfer = set_transfer, rel_eq_onp = rel_eq_onp, pred_transfer = pred_transfer, pred_True = pred_True, pred_map = pred_map, pred_rel = pred_rel, pred_mono_strong0 = pred_mono_strong0, pred_mono_strong = pred_mono_strong, pred_mono = pred_mono, pred_cong0 = pred_cong0, pred_cong = pred_cong, pred_cong_simp = pred_cong_simp}; fun map_facts f { bd_Card_order, bd_Cinfinite, bd_Cnotzero, collect_set_map, in_bd, in_cong, in_mono, in_rel, inj_map, inj_map_strong, map_comp, map_cong, map_cong_simp, map_cong_pred, map_id, map_ident0, map_ident, map_ident_strong, map_transfer, rel_eq, rel_flip, set_map, rel_cong0, rel_cong, rel_cong_simp, rel_map, rel_mono, rel_mono_strong0, rel_mono_strong, rel_transfer, rel_Grp, rel_conversep, rel_OO, rel_refl, rel_refl_strong, rel_reflp, rel_symp, rel_transp, set_transfer, rel_eq_onp, pred_transfer, pred_True, pred_map, pred_rel, pred_mono_strong0, pred_mono_strong, pred_mono, pred_cong0, pred_cong, pred_cong_simp} = {bd_Card_order = f bd_Card_order, bd_Cinfinite = f bd_Cinfinite, bd_Cnotzero = f bd_Cnotzero, collect_set_map = Lazy.map f collect_set_map, in_bd = Lazy.map f in_bd, in_cong = Lazy.map f in_cong, in_mono = Lazy.map f in_mono, in_rel = Lazy.map f in_rel, inj_map = Lazy.map f inj_map, inj_map_strong = Lazy.map f inj_map_strong, map_comp = Lazy.map f map_comp, map_cong = Lazy.map f map_cong, map_cong_simp = Lazy.map f map_cong_simp, map_cong_pred = Lazy.map f map_cong_pred, map_id = Lazy.map f map_id, map_ident0 = Lazy.map f map_ident0, map_ident = Lazy.map f map_ident, map_ident_strong = Lazy.map f map_ident_strong, map_transfer = Lazy.map f map_transfer, rel_eq = Lazy.map f rel_eq, rel_flip = Lazy.map f rel_flip, set_map = map (Lazy.map f) set_map, rel_cong0 = Lazy.map f rel_cong0, rel_cong = Lazy.map f rel_cong, rel_cong_simp = Lazy.map f rel_cong_simp, rel_map = Lazy.map (map f) rel_map, rel_mono = Lazy.map f rel_mono, rel_mono_strong0 = Lazy.map f rel_mono_strong0, rel_mono_strong = Lazy.map f rel_mono_strong, rel_transfer = Lazy.map f rel_transfer, rel_Grp = Lazy.map f rel_Grp, rel_conversep = Lazy.map f rel_conversep, rel_OO = Lazy.map f rel_OO, rel_refl = Lazy.map f rel_refl, rel_refl_strong = Lazy.map f rel_refl_strong, rel_reflp = Lazy.map f rel_reflp, rel_symp = Lazy.map f rel_symp, rel_transp = Lazy.map f rel_transp, set_transfer = Lazy.map (map f) set_transfer, rel_eq_onp = Lazy.map f rel_eq_onp, pred_transfer = Lazy.map f pred_transfer, pred_True = Lazy.map f pred_True, pred_map = Lazy.map f pred_map, pred_rel = Lazy.map f pred_rel, pred_mono_strong0 = Lazy.map f pred_mono_strong0, pred_mono_strong = Lazy.map f pred_mono_strong, pred_mono = Lazy.map f pred_mono, pred_cong0 = Lazy.map f pred_cong0, pred_cong = Lazy.map f pred_cong, pred_cong_simp = Lazy.map f pred_cong_simp}; val morph_facts = map_facts o Morphism.thm; type nonemptiness_witness = { I: int list, wit: term, prop: thm list }; fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop}; fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop}; fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi); datatype bnf = BNF of { name: binding, T: typ, live: int, lives: typ list, (*source type variables of map*) lives': typ list, (*target type variables of map*) dead: int, deads: typ list, map: term, sets: term list, bd: term, axioms: axioms, defs: defs, facts: facts, nwits: int, wits: nonemptiness_witness list, rel: term, pred: term }; (* getters *) fun rep_bnf (BNF bnf) = bnf; val name_of_bnf = #name o rep_bnf; val T_of_bnf = #T o rep_bnf; fun mk_T_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end; val live_of_bnf = #live o rep_bnf; val lives_of_bnf = #lives o rep_bnf; val dead_of_bnf = #dead o rep_bnf; val deads_of_bnf = #deads o rep_bnf; val axioms_of_bnf = #axioms o rep_bnf; val facts_of_bnf = #facts o rep_bnf; val nwits_of_bnf = #nwits o rep_bnf; val wits_of_bnf = #wits o rep_bnf; fun flatten_type_args_of_bnf bnf dead_x xs = let val Type (_, Ts) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val deads = deads_of_bnf bnf; in permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs) end; (*terms*) val map_of_bnf = #map o rep_bnf; val sets_of_bnf = #sets o rep_bnf; fun mk_map_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep) end; fun mk_sets_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; in map2 (fn (Ds, Ts) => Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep) end; val bd_of_bnf = #bd o rep_bnf; fun mk_bd_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end; fun mk_wits_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep); in map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits end; val rel_of_bnf = #rel o rep_bnf; fun mk_rel_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep) end; val pred_of_bnf = #pred o rep_bnf; fun mk_pred_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#pred bnf_rep) end; (*thms*) val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf; val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf; val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf; val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf; val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf; val bd_regularCard_of_bnf = #bd_regularCard o #axioms o rep_bnf; val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf; val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf; val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf; val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf; val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf; val inj_map_of_bnf = Lazy.force o #inj_map o #facts o rep_bnf; val inj_map_strong_of_bnf = Lazy.force o #inj_map_strong o #facts o rep_bnf; val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf; val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf; val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf; val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf; val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf; val map_cong_pred_of_bnf = Lazy.force o #map_cong_pred o #facts o rep_bnf; val map_cong_simp_of_bnf = Lazy.force o #map_cong_simp o #facts o rep_bnf; val map_def_of_bnf = #map_def o #defs o rep_bnf; val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf; val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf; val map_ident0_of_bnf = Lazy.force o #map_ident0 o #facts o rep_bnf; val map_ident_of_bnf = Lazy.force o #map_ident o #facts o rep_bnf; val map_ident_strong_of_bnf = Lazy.force o #map_ident_strong o #facts o rep_bnf; val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf; val rel_eq_onp_of_bnf = Lazy.force o #rel_eq_onp o #facts o rep_bnf; val pred_def_of_bnf = #pred_def o #defs o rep_bnf; val pred_map_of_bnf = Lazy.force o #pred_map o #facts o rep_bnf; val pred_mono_strong0_of_bnf = Lazy.force o #pred_mono_strong0 o #facts o rep_bnf; val pred_mono_strong_of_bnf = Lazy.force o #pred_mono_strong o #facts o rep_bnf; val pred_mono_of_bnf = Lazy.force o #pred_mono o #facts o rep_bnf; val pred_cong0_of_bnf = Lazy.force o #pred_cong0 o #facts o rep_bnf; val pred_cong_of_bnf = Lazy.force o #pred_cong o #facts o rep_bnf; val pred_cong_simp_of_bnf = Lazy.force o #pred_cong_simp o #facts o rep_bnf; val pred_rel_of_bnf = Lazy.force o #pred_rel o #facts o rep_bnf; val pred_set_of_bnf = #pred_set o #axioms o rep_bnf; val pred_transfer_of_bnf = Lazy.force o #pred_transfer o #facts o rep_bnf; val pred_True_of_bnf = Lazy.force o #pred_True o #facts o rep_bnf; val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf; val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf; val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf; val rel_cong0_of_bnf = Lazy.force o #rel_cong0 o #facts o rep_bnf; val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf; val rel_cong_simp_of_bnf = Lazy.force o #rel_cong_simp o #facts o rep_bnf; val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf; val rel_def_of_bnf = #rel_def o #defs o rep_bnf; val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf; val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf; val rel_map_of_bnf = Lazy.force o #rel_map o #facts o rep_bnf; val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf; val rel_mono_strong0_of_bnf = Lazy.force o #rel_mono_strong0 o #facts o rep_bnf; val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf; val rel_refl_of_bnf = Lazy.force o #rel_refl o #facts o rep_bnf; val rel_refl_strong_of_bnf = Lazy.force o #rel_refl_strong o #facts o rep_bnf; val rel_reflp_of_bnf = Lazy.force o #rel_reflp o #facts o rep_bnf; val rel_symp_of_bnf = Lazy.force o #rel_symp o #facts o rep_bnf; val rel_transfer_of_bnf = Lazy.force o #rel_transfer o #facts o rep_bnf; val rel_transp_of_bnf = Lazy.force o #rel_transp o #facts o rep_bnf; val set_bd_of_bnf = #set_bd o #axioms o rep_bnf; val set_defs_of_bnf = #set_defs o #defs o rep_bnf; val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf; val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf; val set_transfer_of_bnf = Lazy.force o #set_transfer o #facts o rep_bnf; val wit_thms_of_bnf = maps #prop o wits_of_bnf; val wit_thmss_of_bnf = map #prop o wits_of_bnf; fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred = BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = length wits, wits = wits, rel = rel, pred = pred}; fun map_bnf f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 (BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = nwits, wits = wits, rel = rel, pred = pred}) = BNF {name = f1 name, T = f2 T, live = f3 live, lives = f4 lives, lives' = f5 lives', dead = f6 dead, deads = f7 deads, map = f8 map, sets = f9 sets, bd = f10 bd, axioms = f11 axioms, defs = f12 defs, facts = f13 facts, nwits = f14 nwits, wits = f15 wits, rel = f16 rel, pred = f17 pred}; fun morph_bnf phi = let val Tphi = Morphism.typ phi; val tphi = Morphism.term phi; in map_bnf (Morphism.binding phi) Tphi I (map Tphi) (map Tphi) I (map Tphi) tphi (map tphi) tphi (morph_axioms phi) (morph_defs phi) (morph_facts phi) I (map (morph_witness phi)) tphi tphi end; fun morph_bnf_defs phi = map_bnf I I I I I I I I I I I (morph_defs phi) I I I I I; fun permute_deads perm = map_bnf I I I I I I perm I I I I I I I I I I; val transfer_bnf = morph_bnf o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = bnf Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun bnf_of_generic context = Option.map (transfer_bnf (Context.theory_of context)) o Symtab.lookup (Data.get context); val bnf_of = bnf_of_generic o Context.Proof; val bnf_of_global = bnf_of_generic o Context.Theory; (* Utilities *) fun normalize_set insts instA set = let val (T, T') = dest_funT (fastype_of set); val A = fst (Term.dest_TVar (HOLogic.dest_setT T')); val params = Term.add_tvar_namesT T []; in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end; fun normalize_rel ctxt instTs instA instB rel = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) rel end handle Type.TYPE_MATCH => error "Bad relator"; fun normalize_pred ctxt instTs instA pred = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of pred, Library.foldr (op -->) (instTs, mk_pred1T instA)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) pred end handle Type.TYPE_MATCH => error "Bad predicator"; fun normalize_wit insts CA As wit = let fun strip_param (Ts, T as Type (\<^type_name>\fun\, [T1, T2])) = if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2) | strip_param x = x; val (Ts, T) = strip_param ([], fastype_of wit); val subst = Term.add_tvar_namesT T [] ~~ insts; fun find y = find_index (fn x => x = y) As; in (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit) end; fun minimize_wits wits = let fun minimize done [] = done | minimize done ((I, wit) :: todo) = if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo) then minimize done todo else minimize ((I, wit) :: done) todo; in minimize [] wits end; fun mk_map live Ts Us t = let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun mk_pred Ts t = let val Type (_, Ts0) = domain_type (body_fun_type (fastype_of t)) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_set = mk_pred; fun mk_rel live Ts Us t = let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts simple_Us build_simple = let fun build (TU as (T, U)) = if exists (curry (op =) T) simple_Ts orelse exists (curry (op =) U) simple_Us then build_simple TU else if T = U andalso not (exists_subtype_in simple_Ts T) andalso not (exists_subtype_in simple_Us U) then const T else (case TU of (Type (s, Ts), Type (s', Us)) => if s = s' then let fun recurse (live, cst0) = let val cst = mk live Ts Us cst0; val TUs' = map dest (fst (strip_typeN live (fastype_of cst))); in Term.list_comb (cst, map build TUs') end; in (case AList.lookup (op =) pre_cst_table s of NONE => (case bnf_of ctxt s of SOME bnf => recurse (live_of_bnf bnf, of_bnf bnf) | NONE => build_simple TU) | SOME entry => recurse entry) end else build_simple TU | _ => build_simple TU); in build end; val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT [(\<^type_name>\set\, (1, \<^term>\image\))]; val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T o append [(\<^type_name>\set\, (1, \<^term>\rel_set\)), (\<^type_name>\fun\, (2, \<^term>\rel_fun\))]; fun build_set ctxt A = let fun build T = Abs (Name.uu, T, if T = A then HOLogic.mk_set A [Bound 0] else (case T of Type (s, Ts) => let val sets = map (mk_set Ts) (sets_of_bnf (the (bnf_of ctxt s))) |> filter (exists_subtype_in [A] o range_type o fastype_of); val set_apps = map (fn set => Term.betapply (set, Bound 0)) sets; fun recurse set_app = let val Type (\<^type_name>\set\, [elemT]) = fastype_of set_app in if elemT = A then set_app else mk_UNION set_app (build elemT) end; in if null set_apps then HOLogic.mk_set A [] else Library.foldl1 mk_union (map recurse set_apps) end | _ => HOLogic.mk_set A [])); in build end; fun map_flattened_map_args ctxt s map_args fs = let val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs; val flat_fs' = map_args flat_fs; in permute_like_unique (op aconv) flat_fs fs flat_fs' end; (* Names *) val mapN = "map"; val setN = "set"; fun mk_setN i = setN ^ nonzero_string_of_int i; val bdN = "bd"; val witN = "wit"; fun mk_witN i = witN ^ nonzero_string_of_int i; val relN = "rel"; val predN = "pred"; val bd_Card_orderN = "bd_Card_order"; val bd_CinfiniteN = "bd_Cinfinite"; val bd_CnotzeroN = "bd_Cnotzero"; val bd_card_orderN = "bd_card_order"; val bd_cinfiniteN = "bd_cinfinite"; val bd_regularCardN = "bd_regularCard"; val collect_set_mapN = "collect_set_map"; val in_bdN = "in_bd"; val in_monoN = "in_mono"; val in_relN = "in_rel"; val inj_mapN = "inj_map"; val inj_map_strongN = "inj_map_strong"; val map_comp0N = "map_comp0"; val map_compN = "map_comp"; val map_cong0N = "map_cong0"; val map_congN = "map_cong"; val map_cong_simpN = "map_cong_simp"; val map_cong_predN = "map_cong_pred"; val map_id0N = "map_id0"; val map_idN = "map_id"; val map_identN = "map_ident"; val map_ident_strongN = "map_ident_strong"; val map_transferN = "map_transfer"; val pred_mono_strong0N = "pred_mono_strong0"; val pred_mono_strongN = "pred_mono_strong"; val pred_monoN = "pred_mono"; val pred_transferN = "pred_transfer"; val pred_TrueN = "pred_True"; val pred_mapN = "pred_map"; val pred_relN = "pred_rel"; val pred_setN = "pred_set"; val pred_congN = "pred_cong"; val pred_cong_simpN = "pred_cong_simp"; val rel_GrpN = "rel_Grp"; val rel_comppN = "rel_compp"; val rel_compp_GrpN = "rel_compp_Grp"; val rel_congN = "rel_cong"; val rel_cong_simpN = "rel_cong_simp"; val rel_conversepN = "rel_conversep"; val rel_eqN = "rel_eq"; val rel_eq_onpN = "rel_eq_onp"; val rel_flipN = "rel_flip"; val rel_mapN = "rel_map"; val rel_monoN = "rel_mono"; val rel_mono_strong0N = "rel_mono_strong0"; val rel_mono_strongN = "rel_mono_strong"; val rel_reflN = "rel_refl"; val rel_refl_strongN = "rel_refl_strong"; val rel_reflpN = "rel_reflp"; val rel_sympN = "rel_symp"; val rel_transferN = "rel_transfer"; val rel_transpN = "rel_transp"; val set_bdN = "set_bd"; val set_map0N = "set_map0"; val set_mapN = "set_map"; val set_transferN = "set_transfer"; datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline; datatype fact_policy = Dont_Note | Note_Some | Note_All; val bnf_internals = Attrib.setup_config_bool \<^binding>\bnf_internals\ (K false); val bnf_timing = Attrib.setup_config_bool \<^binding>\bnf_timing\ (K false); fun user_policy policy ctxt = if Config.get ctxt bnf_internals then Note_All else policy; val smart_max_inline_term_size = 25; (*FUDGE*) fun note_bnf_thms fact_policy qualify0 bnf_b bnf lthy = let val axioms = axioms_of_bnf bnf; val facts = facts_of_bnf bnf; val wits = wits_of_bnf bnf; val qualify = let val qs = Binding.path_of bnf_b; in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify0 end; fun note_if_note_all (noted0, lthy0) = let val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits); val notes = [(bd_card_orderN, [#bd_card_order axioms]), (bd_cinfiniteN, [#bd_cinfinite axioms]), (bd_regularCardN, [#bd_regularCard axioms]), (bd_Card_orderN, [#bd_Card_order facts]), (bd_CinfiniteN, [#bd_Cinfinite facts]), (bd_CnotzeroN, [#bd_Cnotzero facts]), (collect_set_mapN, [Lazy.force (#collect_set_map facts)]), (in_bdN, [Lazy.force (#in_bd facts)]), (in_monoN, [Lazy.force (#in_mono facts)]), (map_comp0N, [#map_comp0 axioms]), (rel_mono_strong0N, [Lazy.force (#rel_mono_strong0 facts)]), (pred_mono_strong0N, [Lazy.force (#pred_mono_strong0 facts)]), (set_map0N, #set_map0 axioms), (set_bdN, #set_bd axioms)] @ (witNs ~~ wit_thmss_of_bnf bnf) |> map (fn (thmN, thms) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; fun note_unless_dont_note (noted0, lthy0) = let val notes = [(in_relN, [Lazy.force (#in_rel facts)], []), (inj_mapN, [Lazy.force (#inj_map facts)], []), (inj_map_strongN, [Lazy.force (#inj_map_strong facts)], []), (map_compN, [Lazy.force (#map_comp facts)], []), (map_cong0N, [#map_cong0 axioms], []), (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs), (map_cong_simpN, [Lazy.force (#map_cong_simp facts)], []), (map_cong_predN, [Lazy.force (#map_cong_pred facts)], []), (map_idN, [Lazy.force (#map_id facts)], []), (map_id0N, [#map_id0 axioms], []), (map_transferN, [Lazy.force (#map_transfer facts)], []), (map_identN, [Lazy.force (#map_ident facts)], []), (map_ident_strongN, [Lazy.force (#map_ident_strong facts)], []), (pred_monoN, [Lazy.force (#pred_mono facts)], mono_attrs), (pred_mono_strongN, [Lazy.force (#pred_mono_strong facts)], []), (pred_congN, [Lazy.force (#pred_cong facts)], fundefcong_attrs), (pred_cong_simpN, [Lazy.force (#pred_cong_simp facts)], []), (pred_mapN, [Lazy.force (#pred_map facts)], []), (pred_relN, [Lazy.force (#pred_rel facts)], []), (pred_transferN, [Lazy.force (#pred_transfer facts)], []), (pred_TrueN, [Lazy.force (#pred_True facts)], []), (pred_setN, [#pred_set axioms], []), (rel_comppN, [Lazy.force (#rel_OO facts)], []), (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []), (rel_conversepN, [Lazy.force (#rel_conversep facts)], []), (rel_eqN, [Lazy.force (#rel_eq facts)], []), (rel_eq_onpN, [Lazy.force (#rel_eq_onp facts)], []), (rel_flipN, [Lazy.force (#rel_flip facts)], []), (rel_GrpN, [Lazy.force (#rel_Grp facts)], []), (rel_mapN, Lazy.force (#rel_map facts), []), (rel_monoN, [Lazy.force (#rel_mono facts)], mono_attrs), (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)], []), (rel_congN, [Lazy.force (#rel_cong facts)], fundefcong_attrs), (rel_cong_simpN, [Lazy.force (#rel_cong_simp facts)], []), (rel_reflN, [Lazy.force (#rel_refl facts)], []), (rel_refl_strongN, [Lazy.force (#rel_refl_strong facts)], []), (rel_reflpN, [Lazy.force (#rel_reflp facts)], []), (rel_sympN, [Lazy.force (#rel_symp facts)], []), (rel_transpN, [Lazy.force (#rel_transp facts)], []), (rel_transferN, [Lazy.force (#rel_transfer facts)], []), (set_mapN, map Lazy.force (#set_map facts), []), (set_transferN, Lazy.force (#set_transfer facts), [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), attrs), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; in ([], lthy) |> fact_policy = Note_All ? note_if_note_all |> fact_policy <> Dont_Note ? note_unless_dont_note |>> (fn [] => bnf | noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun note_bnf_defs bnf lthy = let fun mk_def_binding cst_of = Thm.def_binding (Binding.qualified_name (dest_Const (cst_of bnf) |> fst)); val notes = [(mk_def_binding map_of_bnf, map_def_of_bnf bnf), (mk_def_binding rel_of_bnf, rel_def_of_bnf bnf), (mk_def_binding pred_of_bnf, pred_def_of_bnf bnf)] @ @{map 2} (pair o mk_def_binding o K) (sets_of_bnf bnf) (set_defs_of_bnf bnf) |> map (fn (b, thm) => ((b, []), [([thm], [])])); in lthy |> Local_Theory.notes notes |>> (fn noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun mk_wit_goals zs bs sets (I, wit) = let val xs = map (nth bs) I; fun wit_goal i = let val z = nth zs i; val set_wit = nth sets i $ Term.list_comb (wit, xs); val concl = HOLogic.mk_Trueprop (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else \<^term>\False\); in fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl)) end; in map wit_goal (0 upto length sets - 1) end; (* Define new BNFs *) fun define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy = let val live = length set_rhss; val def_qualify = Binding.qualify false (Binding.name_of bnf_b); fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b; fun maybe_define user_specified (b, rhs) lthy = let val inline = (user_specified orelse fact_policy = Dont_Note) andalso (case const_policy of Dont_Inline => false | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_term_size | Do_Inline => true); in if inline then ((rhs, Drule.reflexive_thm), lthy) else let val b = b () in apfst (apsnd snd) ((if internal then Local_Theory.define_internal else Local_Theory.define) ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), rhs)) lthy) end end; val map_bind_def = (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b), map_rhs); val set_binds_defs = let fun set_name i get_b = (case try (nth set_bs) (i - 1) of SOME b => if Binding.is_empty b then get_b else K b | NONE => get_b) #> def_qualify; val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)] else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live); in bs ~~ set_rhss end; val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs); val (((bnf_map_term, raw_map_def), (bnf_set_terms, raw_set_defs)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> maybe_define true map_bind_def ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val ((bnf_bd_term, raw_bd_def), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> maybe_define true bd_bind_def ||> `Local_Theory.end_nested; val phi' = Proof_Context.export_morphism lthy_old lthy; val bnf_map_def = Morphism.thm phi raw_map_def; val bnf_set_defs = map (Morphism.thm phi) raw_set_defs; val bnf_bd_def = Morphism.thm phi' raw_bd_def; val bnf_map = Morphism.term phi bnf_map_term; (*TODO: handle errors*) (*simple shape analysis of a map function*) val ((alphas, betas), (Calpha, _)) = fastype_of bnf_map |> strip_typeN live |>> map_split dest_funT ||> dest_funT handle TYPE _ => error "Bad map function"; val Calpha_params = map TVar (Term.add_tvarsT Calpha []); val bnf_T = Morphism.typ phi T_rhs; val bad_args = Term.add_tfreesT bnf_T []; val _ = null bad_args orelse error ("Locally fixed type arguments " ^ commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args)); val bnf_sets = map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms); val bnf_bd = Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params) (Morphism.term phi' bnf_bd_term); (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*) val deads = (case Ds_opt of NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map [])) | SOME Ds => map (Morphism.typ phi) Ds); (*TODO: further checks of type of bnf_map*) (*TODO: check types of bnf_sets*) (*TODO: check type of bnf_bd*) (*TODO: check type of bnf_rel*) fun mk_bnf_map Ds As' Bs' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map; fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As')); fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As')); val (((As, Bs), unsorted_Ds), names_lthy) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees (length deads); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val RTs = map2 (curry HOLogic.mk_prodT) As Bs; val pred2RTs = map2 mk_pred2T As Bs; val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst; val CA = mk_bnf_T Ds As Calpha; val CR = mk_bnf_T Ds RTs Calpha; val setRs = @{map 3} (fn R => fn T => fn U => HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_case_prod R) Rs As Bs; (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*) val rel_spec = let val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs); val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs); val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR; in mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2) |> fold_rev Term.absfree Rs' end; val rel_rhs = the_default rel_spec rel_rhs_opt; val rel_bind_def = (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b), rel_rhs); val pred_spec = if live = 0 then Term.absdummy (mk_bnf_T Ds As Calpha) \<^term>\True\ else let val sets = map (mk_bnf_t Ds As) bnf_sets; val argTs = map mk_pred1T As; val T = mk_bnf_T Ds As Calpha; val ((Ps, Ps'), x) = lthy |> mk_Frees' "P" argTs ||>> yield_singleton (mk_Frees "x") T |> fst; val conjs = map2 (fn set => fn P => mk_Ball (set $ x) P) sets Ps; in fold_rev Term.absfree Ps' (Term.absfree (dest_Free x) (Library.foldr1 HOLogic.mk_conj conjs)) end; val pred_rhs = the_default pred_spec pred_rhs_opt; val pred_bind_def = (fn () => def_qualify (if Binding.is_empty pred_b then mk_prefix_binding predN else pred_b), pred_rhs); val wit_rhss = if null wit_rhss then [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As, map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $ Const (\<^const_name>\undefined\, CA))] else wit_rhss; val nwits = length wit_rhss; val wit_binds_defs = let val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)] else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits); in bs ~~ wit_rhss end; val ((((bnf_rel_term, raw_rel_def), (bnf_pred_term, raw_pred_def)), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> maybe_define (is_some rel_rhs_opt) rel_bind_def ||>> maybe_define (is_some pred_rhs_opt) pred_bind_def ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val bnf_rel_def = Morphism.thm phi raw_rel_def; val bnf_rel = Morphism.term phi bnf_rel_term; fun mk_bnf_rel Ds As' Bs' = normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha) bnf_rel; val bnf_pred_def = Morphism.thm phi raw_pred_def; val bnf_pred = Morphism.term phi bnf_pred_term; fun mk_bnf_pred Ds As' = normalize_pred lthy (map mk_pred1T As') (mk_bnf_T Ds As' Calpha) bnf_pred; val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs; val bnf_wits = map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms; fun mk_rel_spec Ds' As' Bs' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) rel_spec; fun mk_pred_spec Ds' As' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As')) pred_spec; in (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_bnf_pred, mk_rel_spec, mk_pred_spec)), lthy) end; fun prepare_def const_policy mk_fact_policy internal qualify prep_typ prep_term Ds_opt map_b rel_b pred_b set_bs (((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt), raw_pred_opt) no_defs_lthy = let val fact_policy = mk_fact_policy no_defs_lthy; val bnf_b = qualify raw_bnf_b; val live = length raw_sets; val T_rhs = prep_typ no_defs_lthy raw_bnf_T; val map_rhs = prep_term no_defs_lthy raw_map; val set_rhss = map (prep_term no_defs_lthy) raw_sets; val bd_rhs = prep_term no_defs_lthy raw_bd; val wit_rhss = map (prep_term no_defs_lthy) raw_wits; val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt; val pred_rhs_opt = Option.map (prep_term no_defs_lthy) raw_pred_opt; fun err T = error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^ " as unnamed BNF"); val (bnf_b, key) = if Binding.is_empty bnf_b then (case T_rhs of Type (C, Ts) => if forall (can dest_TFree) Ts andalso not (has_duplicates (op =) Ts) then (Binding.qualified_name C, C) else err T_rhs | T => err T) else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b); val (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, _, mk_rel_spec, mk_pred_spec)), lthy) = define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy; val dead = length deads; val (((((((As', Bs'), Cs), unsorted_Ds), Es), B1Ts), B2Ts), (Ts, T)) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees dead ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||> fst o mk_TFrees 1 ||> the_single ||> `(replicate live); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val mk_bnf_map = mk_bnf_map_Ds Ds; val mk_bnf_t = mk_bnf_t_Ds Ds; val mk_bnf_T = mk_bnf_T_Ds Ds; val pred1PTs = map mk_pred1T As'; val pred1QTs = map mk_pred1T Bs'; val pred2RTs = map2 mk_pred2T As' Bs'; val pred2RTsAsCs = map2 mk_pred2T As' Cs; val pred2RTsBsCs = map2 mk_pred2T Bs' Cs; val pred2RTsBsEs = map2 mk_pred2T Bs' Es; val pred2RTsCsBs = map2 mk_pred2T Cs Bs'; val pred2RTsCsEs = map2 mk_pred2T Cs Es; val pred2RT's = map2 mk_pred2T Bs' As'; val self_pred2RTs = map2 mk_pred2T As' As'; val transfer_domRTs = map2 mk_pred2T As' B1Ts; val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts; val CA' = mk_bnf_T As' Calpha; val CB' = mk_bnf_T Bs' Calpha; val CC' = mk_bnf_T Cs Calpha; val CE' = mk_bnf_T Es Calpha; val CB1 = mk_bnf_T B1Ts Calpha; val CB2 = mk_bnf_T B2Ts Calpha; val bnf_map_AsAs = mk_bnf_map As' As'; val bnf_map_AsBs = mk_bnf_map As' Bs'; val bnf_map_AsCs = mk_bnf_map As' Cs; val bnf_map_BsCs = mk_bnf_map Bs' Cs; val bnf_sets_As = map (mk_bnf_t As') bnf_sets; val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets; val bnf_bd_As = mk_bnf_t As' bnf_bd; fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel; fun mk_bnf_pred PTs CA = normalize_pred lthy PTs CA bnf_pred; val ((((((((((((((((((((((((((fs, fs'), gs), hs), is), x), x'), y), y'), zs), zs'), ys), As), As_copy), bs), (Ps, Ps')), Ps_copy), Qs), Rs), Rs_copy), Ss), S_AsCs), S_CsBs), S_BsEs), transfer_domRs), transfer_ranRs), _) = lthy |> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs) ||>> mk_Frees "h" (map2 (curry op -->) As' Ts) ||>> mk_Frees "i" (map2 (curry op -->) As' Cs) ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "y") CB' ||>> yield_singleton (mk_Frees "y") CB' ||>> mk_Frees "z" As' ||>> mk_Frees "z" As' ||>> mk_Frees "y" Bs' ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "b" As' ||>> mk_Frees' "P" pred1PTs ||>> mk_Frees "P" pred1PTs ||>> mk_Frees "Q" pred1QTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "S" pred2RTsBsCs ||>> mk_Frees "S" pred2RTsAsCs ||>> mk_Frees "S" pred2RTsCsBs ||>> mk_Frees "S" pred2RTsBsEs ||>> mk_Frees "R" transfer_domRTs ||>> mk_Frees "S" transfer_ranRTs; val fs_copy = map2 (retype_const_or_free o fastype_of) fs gs; val x_copy = retype_const_or_free CA' y'; val y_copy = retype_const_or_free CB' x'; val rel = mk_bnf_rel pred2RTs CA' CB'; val pred = mk_bnf_pred pred1PTs CA'; val pred' = mk_bnf_pred pred1QTs CB'; val relCsEs = mk_bnf_rel pred2RTsCsEs CC' CE'; val relAsAs = mk_bnf_rel self_pred2RTs CA' CA'; val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits; val map_id0_goal = let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA') end; val map_comp0_goal = let val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs); val comp_bnf_map_app = HOLogic.mk_comp (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs)); in fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app)) end; fun mk_map_cong_prem mk_implies x z set f f_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (f $ z, f_copy $ z))); val map_cong0_goal = let val prems = @{map 4} (mk_map_cong_prem Logic.mk_implies x) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x); in fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq)) end; val set_map0s_goal = let fun mk_goal setA setB f = let val set_comp_map = HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs)); val image_comp_set = HOLogic.mk_comp (mk_image f, setA); in fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set)) end; in @{map 3} mk_goal bnf_sets_As bnf_sets_Bs fs end; val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As); val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As); val regularCard_bd_goal = HOLogic.mk_Trueprop (mk_regularCard bnf_bd_As); val set_bds_goal = let fun mk_goal set = Logic.all x (HOLogic.mk_Trueprop (mk_ordLess (mk_card_of (set $ x)) bnf_bd_As)); in map mk_goal bnf_sets_As end; val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC'; val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC'; val relCsBs = mk_bnf_rel pred2RTsCsBs CC' CB'; val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss); val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss)); val le_rel_OO_goal = fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs)); val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs), Term.list_comb (mk_rel_spec Ds As' Bs', Rs))); val pred_set_goal = fold_rev Logic.all Ps (mk_Trueprop_eq (Term.list_comb (pred, Ps), Term.list_comb (mk_pred_spec Ds As', Ps))); val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal card_order_bd_goal cinfinite_bd_goal regularCard_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal pred_set_goal; val mk_wit_goals = mk_wit_goals bs zs bnf_sets_As; fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs; val wit_goalss = (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As); fun after_qed mk_wit_thms thms lthy = let val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms); val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]}; val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order]; val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero}; fun mk_collect_set_map () = let val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T; val collect_map = HOLogic.mk_comp (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT, Term.list_comb (mk_bnf_map As' Ts, hs)); val image_collect = mk_collect (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As) defT; (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*) val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_collect_set_map_tac ctxt (#set_map0 axioms)) |> Thm.close_derivation \<^here> end; val collect_set_map = Lazy.lazy mk_collect_set_map; fun mk_in_mono () = let val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy; val in_mono_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA')))); in Goal.prove_sorry lthy [] [] in_mono_goal (fn {context = ctxt, prems = _} => mk_in_mono_tac ctxt live) |> Thm.close_derivation \<^here> end; val in_mono = Lazy.lazy mk_in_mono; fun mk_in_cong () = let val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy; val in_cong_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_cong, mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))); in Goal.prove_sorry lthy [] [] in_cong_goal (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val in_cong = Lazy.lazy mk_in_cong; val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms)); val map_ident0 = Lazy.lazy (fn () => mk_map_ident lthy (#map_id0 axioms)); val map_ident = Lazy.lazy (fn () => mk_map_ident lthy (Lazy.force map_id)); val map_ident_strong = Lazy.lazy (fn () => mk_map_ident_strong lthy (#map_cong0 axioms) (Lazy.force map_id)); val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms)); fun mk_map_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_map_cong_prem mk_implies x_copy) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies (prem0 :: prems, eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt @{thms simp_implies_def} THEN mk_map_cong_tac ctxt (#map_cong0 axioms)) |> Thm.close_derivation \<^here> end; val map_cong = Lazy.lazy (mk_map_cong Logic.mk_implies); val map_cong_simp = Lazy.lazy (mk_map_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_inj_map () = let val prems = map (HOLogic.mk_Trueprop o mk_inj) fs; val concl = HOLogic.mk_Trueprop (mk_inj (Term.list_comb (bnf_map_AsBs, fs))); val goal = fold_rev Logic.all fs (Logic.list_implies (prems, concl)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_tac ctxt live (Lazy.force map_id) (Lazy.force map_comp) (#map_cong0 axioms) (Lazy.force map_cong)) |> Thm.close_derivation \<^here> end; val inj_map = Lazy.lazy mk_inj_map; val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms); val wit_thms = if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms; fun mk_in_bd () = let val bdT = fst (dest_relT (fastype_of bnf_bd_As)); val bdTs = replicate live bdT; val bd_bnfT = mk_bnf_T bdTs Calpha; val surj_imp_ordLeq_inst = (if live = 0 then TrueI else let val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As'; val funTs = map (fn T => bdT --> T) ranTs; val ran_bnfT = mk_bnf_T ranTs Calpha; val (revTs, Ts) = `rev (bd_bnfT :: funTs); val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; val tinst = fold (fn T => fn t => HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs) (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, map Bound (live - 1 downto 0)) $ Bound live)); val cts = [NONE, SOME (Thm.cterm_of lthy tinst)]; in Thm.instantiate' cTs cts @{thm surj_imp_ordLeq} end); val bd = mk_cexp (if live = 0 then ctwo else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT))); val in_bd_goal = fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd)); val weak_set_bds = map (fn thm => @{thm ordLess_imp_ordLeq} OF [thm]) (#set_bd axioms); in Goal.prove_sorry lthy [] [] in_bd_goal (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms) (map Lazy.force set_map) weak_set_bds (#bd_card_order axioms) bd_Card_order bd_Cinfinite bd_Cnotzero) |> Thm.close_derivation \<^here> end; val in_bd = Lazy.lazy mk_in_bd; val rel_OO_Grp = #rel_OO_Grp axioms; val rel_OO_Grps = no_refl [rel_OO_Grp]; fun mk_rel_Grp () = let val lhs = Term.list_comb (rel, map2 mk_Grp As fs); val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs)); val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_Grp = Lazy.lazy mk_rel_Grp; fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy; fun mk_rel_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy))); fun mk_rel_mono () = let val mono_prems = mk_rel_prems mk_leq; val mono_concl = mk_rel_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_tac ctxt rel_OO_Grps (Lazy.force in_mono)) |> Thm.close_derivation \<^here> end; fun mk_rel_cong0 () = let val cong_prems = mk_rel_prems (curry HOLogic.mk_eq); val cong_concl = mk_rel_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val rel_mono = Lazy.lazy mk_rel_mono; val rel_cong0 = Lazy.lazy mk_rel_cong0; fun mk_rel_eq () = Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'), HOLogic.eq_const CA')) (fn {context = ctxt, prems = _} => mk_rel_eq_tac ctxt live (Lazy.force rel_Grp) (Lazy.force rel_cong0) (#map_id0 axioms)) |> Thm.close_derivation \<^here>; val rel_eq = Lazy.lazy mk_rel_eq; fun mk_rel_conversep () = let val relBsAs = mk_bnf_rel pred2RT's CB' CA'; val lhs = Term.list_comb (relBsAs, map mk_conversep Rs); val rhs = mk_conversep (Term.list_comb (rel, Rs)); val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs)); val le_thm = Goal.prove_sorry lthy [] [] le_goal (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_conversep_tac ctxt le_thm (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val rel_conversep = Lazy.lazy mk_rel_conversep; fun mk_rel_OO () = Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs))) (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]); val rel_OO = Lazy.lazy mk_rel_OO; fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}; val in_rel = Lazy.lazy mk_in_rel; fun mk_rel_flip () = unfold_thms lthy @{thms conversep_iff} (Lazy.force rel_conversep RS @{thm predicate2_eqD}); val rel_flip = Lazy.lazy mk_rel_flip; fun mk_rel_mono_strong0 () = let fun mk_prem setA setB R S a b = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (mk_Ball (setB $ y) (Term.absfree (dest_Free b) (HOLogic.mk_imp (R $ a $ b, S $ a $ b)))))); val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: @{map 6} mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys; val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_strong0_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_mono_strong0 = Lazy.lazy mk_rel_mono_strong0; val rel_mono_strong = Lazy.map (Object_Logic.rulify lthy) rel_mono_strong0; fun mk_rel_cong_prem mk_implies x x' z z' set set' R R_copy = Logic.all z (Logic.all z' (mk_implies (mk_Trueprop_mem (z, set $ x), mk_implies (mk_Trueprop_mem (z', set' $ x'), mk_Trueprop_eq (R $ z $ z', R_copy $ z $ z'))))); fun mk_rel_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prem1 = mk_Trueprop_eq (y, y_copy); val prems = @{map 6} (mk_rel_cong_prem mk_implies x_copy y_copy) zs ys bnf_sets_As bnf_sets_Bs Rs Rs_copy; val eq = mk_Trueprop_eq (Term.list_comb (rel, Rs) $ x $ y, Term.list_comb (rel, Rs_copy) $ x_copy $ y_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prem1 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prem1 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 2 prems) (Lazy.force rel_mono_strong))) |> Thm.close_derivation \<^here> end; val rel_cong = Lazy.lazy (mk_rel_cong Logic.mk_implies); val rel_cong_simp = Lazy.lazy (mk_rel_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_pred_prems f = map2 (HOLogic.mk_Trueprop oo f) Ps Ps_copy; fun mk_pred_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (pred, Ps), Term.list_comb (pred, Ps_copy))); fun mk_pred_cong0 () = let val cong_prems = mk_pred_prems (curry HOLogic.mk_eq); val cong_concl = mk_pred_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val pred_cong0 = Lazy.lazy mk_pred_cong0; fun mk_rel_eq_onp () = let val lhs = Term.list_comb (relAsAs, map mk_eq_onp Ps); val rhs = mk_eq_onp (Term.list_comb (pred, Ps)); in Goal.prove_sorry lthy (map fst Ps') [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_rel_eq_onp_tac ctxt (#pred_set axioms) (#map_id0 axioms) (Lazy.force rel_Grp)) |> Thm.close_derivation \<^here> end; val rel_eq_onp = Lazy.lazy mk_rel_eq_onp; val pred_rel = Lazy.map (fn thm => thm RS sym RS @{thm eq_onp_eqD}) rel_eq_onp; fun mk_pred_mono_strong0 () = let fun mk_prem setA P Q a = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (HOLogic.mk_imp (P $ a, Q $ a)))); val prems = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps) $ x) :: @{map 4} mk_prem bnf_sets_As Ps Ps_copy zs; val concl = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps_copy) $ x); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: Ps @ Ps_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_strong0_tac ctxt (Lazy.force pred_rel) (Lazy.force rel_mono_strong0)) |> Thm.close_derivation \<^here> end; val pred_mono_strong0 = Lazy.lazy mk_pred_mono_strong0; val pred_mono_strong = Lazy.map (Object_Logic.rulify lthy) pred_mono_strong0; fun mk_pred_mono () = let val mono_prems = mk_pred_prems mk_leq; val mono_concl = mk_pred_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_tac ctxt (Lazy.force rel_eq_onp) (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val pred_mono = Lazy.lazy mk_pred_mono; fun mk_pred_cong_prem mk_implies x z set P P_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (P $ z, P_copy $ z))); fun mk_pred_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_pred_cong_prem mk_implies x_copy) zs bnf_sets_As Ps Ps_copy; val eq = mk_Trueprop_eq (Term.list_comb (pred, Ps) $ x, Term.list_comb (pred, Ps_copy) $ x_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 1 prems) (Lazy.force pred_mono_strong))) |> Thm.close_derivation \<^here> end; val pred_cong = Lazy.lazy (mk_pred_cong Logic.mk_implies); val pred_cong_simp = Lazy.lazy (mk_pred_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_map_cong_pred () = let val prem0 = mk_Trueprop_eq (x, x_copy); fun mk_eq f g z = Term.absfree (dest_Free z) (HOLogic.mk_eq (f $ z, g $ z)); val prem = HOLogic.mk_Trueprop (Term.list_comb (pred, @{map 3} mk_eq fs fs_copy zs) $ x_copy); val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies ([prem0, prem], eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [#pred_set axioms] THEN HEADGOAL (EVERY' [REPEAT_DETERM o etac ctxt conjE, etac ctxt (Lazy.force map_cong) THEN_ALL_NEW (etac ctxt @{thm bspec} THEN' assume_tac ctxt)])) |> Thm.close_derivation \<^here> end; val map_cong_pred = Lazy.lazy mk_map_cong_pred; fun mk_rel_map () = let fun mk_goal lhs rhs = fold_rev Logic.all ([x, y] @ S_CsBs @ S_AsCs @ is @ gs) (mk_Trueprop_eq (lhs, rhs)); val lhss = [Term.list_comb (relCsBs, S_CsBs) $ (Term.list_comb (bnf_map_AsCs, is) $ x) $ y, Term.list_comb (relAsCs, S_AsCs) $ x $ (Term.list_comb (bnf_map_BsCs, gs) $ y)]; val rhss = [Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p f (HOLogic.id_const T) $ P) is S_CsBs Bs') $ x $ y, Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p (HOLogic.id_const T) f $ P) gs S_AsCs As') $ x $ y]; val goals = map2 mk_goal lhss rhss; in goals |> map (fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_map0_tac ctxt live (Lazy.force rel_OO) (Lazy.force rel_conversep) (Lazy.force rel_Grp) (Lazy.force map_id))) |> map (unfold_thms lthy @{thms vimage2p_def[of id, simplified id_apply] vimage2p_def[of _ id, simplified id_apply]}) |> map (Thm.close_derivation \<^here>) end; val rel_map = Lazy.lazy mk_rel_map; fun mk_rel_refl () = @{thm ge_eq_refl[OF ord_eq_le_trans]} OF [Lazy.force rel_eq RS sym, Lazy.force rel_mono OF (replicate live @{thm refl_ge_eq})]; val rel_refl = Lazy.lazy mk_rel_refl; fun mk_rel_refl_strong () = (rule_by_tactic lthy (ALLGOALS (Object_Logic.full_atomize_tac lthy)) ((Lazy.force rel_eq RS @{thm predicate2_eqD}) RS @{thm iffD2[OF _ refl]} RS Lazy.force rel_mono_strong)) OF (replicate live @{thm diag_imp_eq_le}) val rel_refl_strong = Lazy.lazy mk_rel_refl_strong; fun mk_rel_preserves mk_prop prop_conv_thm thm () = let val Rs = map2 retype_const_or_free self_pred2RTs Rs; val prems = map (HOLogic.mk_Trueprop o mk_prop) Rs; val goal = HOLogic.mk_Trueprop (mk_prop (Term.list_comb (relAsAs, Rs))); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [prop_conv_thm] THEN HEADGOAL (rtac ctxt (Lazy.force thm RS sym RS @{thm ord_eq_le_trans}) THEN' rtac ctxt (Lazy.force rel_mono) THEN_ALL_NEW assume_tac ctxt)) |> Thm.close_derivation \<^here> end; val rel_reflp = Lazy.lazy (mk_rel_preserves mk_reflp @{thm reflp_eq} rel_eq); val rel_symp = Lazy.lazy (mk_rel_preserves mk_symp @{thm symp_conversep} rel_conversep); val rel_transp = Lazy.lazy (mk_rel_preserves mk_transp @{thm transp_relcompp} rel_OO); fun mk_pred_True () = let val lhs = Term.list_comb (pred, map (fn T => absdummy T \<^term>\True\) As'); val rhs = absdummy CA' \<^term>\True\; val goal = mk_Trueprop_eq (lhs, rhs); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (EVERY' (map (rtac ctxt) [ext, Lazy.force pred_rel RS trans, Lazy.force rel_cong0 RS fun_cong RS fun_cong RS trans OF replicate live @{thm eq_onp_True}, Lazy.force rel_eq RS fun_cong RS fun_cong RS trans, @{thm eqTrueI[OF refl]}]))) |> Thm.close_derivation \<^here> end; val pred_True = Lazy.lazy mk_pred_True; fun mk_pred_map () = let val lhs = Term.list_comb (pred', Qs) $ (Term.list_comb (bnf_map_AsBs, fs) $ x); val rhs = Term.list_comb (pred, @{map 2} (curry HOLogic.mk_comp) Qs fs) $ x; val goal = mk_Trueprop_eq (lhs, rhs); val vars = Variable.add_free_names lthy goal []; val pred_set = #pred_set axioms RS fun_cong RS sym; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt (pred_set RSN (2, pred_set RSN (2, box_equals)))) THEN unfold_thms_tac ctxt (@{thms Ball_image_comp ball_empty} @ map Lazy.force set_map) THEN HEADGOAL (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; val pred_map = Lazy.lazy mk_pred_map; fun mk_map_transfer () = let val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs; val rel = mk_rel_fun (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs)) (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs)); val concl = HOLogic.mk_Trueprop (fold_rev mk_rel_fun rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl) (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono) (Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp)) |> Thm.close_derivation \<^here> end; val map_transfer = Lazy.lazy mk_map_transfer; fun mk_pred_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map (fn T => mk_rel_fun T iff) Rs; val prem_elems = mk_rel_fun (Term.list_comb (rel, Rs)) iff; val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ pred $ pred'); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_pred_transfer_tac ctxt live (Lazy.force in_rel) (Lazy.force pred_map) (Lazy.force pred_cong)) |> Thm.close_derivation \<^here> end; val pred_transfer = Lazy.lazy mk_pred_transfer; fun mk_rel_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map2 (fn T1 => fn T2 => mk_rel_fun T1 (mk_rel_fun T2 iff)) S_AsCs S_BsEs; val prem_elems = mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsAsCs CA' CC', S_AsCs)) (mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsBsEs CB' CE', S_BsEs)) iff); val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ rel $ relCsEs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_transfer_tac ctxt (Lazy.force in_rel) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val rel_transfer = Lazy.lazy mk_rel_transfer; fun mk_set_transfer () = let val rel_sets = map2 (fn A => fn B => mk_rel 1 [A] [B] \<^term>\rel_set\) As' Bs'; val rel_Rs = Term.list_comb (rel, Rs); val goals = @{map 4} (fn R => fn rel_set => fn setA => fn setB => HOLogic.mk_Trueprop (mk_rel_fun rel_Rs (rel_set $ R) $ setA $ setB)) Rs rel_sets bnf_sets_As bnf_sets_Bs; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_transfer_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_transfer = Lazy.lazy mk_set_transfer; fun mk_inj_map_strong () = let val assms = @{map 5} (fn setA => fn z => fn f => fn z' => fn f' => fold_rev Logic.all [z, z'] (Logic.mk_implies (mk_Trueprop_mem (z, setA $ x), Logic.mk_implies (mk_Trueprop_mem (z', setA $ x'), Logic.mk_implies (mk_Trueprop_eq (f $ z, f' $ z'), mk_Trueprop_eq (z, z')))))) bnf_sets_As zs fs zs' fs'; val concl = Logic.mk_implies (mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs') $ x'), mk_Trueprop_eq (x, x')); val goal = fold_rev Logic.all (x :: x' :: fs @ fs') (fold_rev (curry Logic.mk_implies) assms concl); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_strong_tac ctxt (Lazy.force rel_eq) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val inj_map_strong = Lazy.lazy mk_inj_map_strong; val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_pred_def; val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp; val wits = map2 mk_witness bnf_wits wit_thms; val bnf_rel = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel; val bnf_pred = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas)) pred; val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts wits bnf_rel bnf_pred; in note_bnf_thms fact_policy qualify bnf_b bnf lthy end; val one_step_defs = no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def, bnf_pred_def]); in (key, goals, wit_goalss, after_qed, lthy, one_step_defs) end; structure BNF_Plugin = Plugin(type T = bnf); fun bnf_interpretation name f = BNF_Plugin.interpretation name (fn bnf => fn lthy => f (transfer_bnf (Proof_Context.theory_of lthy) bnf) lthy); val interpret_bnf = BNF_Plugin.data; fun register_bnf_raw key bnf = - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf))); fun register_bnf plugins key bnf = register_bnf_raw key bnf #> interpret_bnf plugins bnf; fun bnf_def const_policy fact_policy internal qualify tacs wit_tac Ds map_b rel_b pred_b set_bs raw_csts = (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) => let fun mk_wits_tac ctxt set_maps = TRYALL Goal.conjunction_tac THEN (case triv_tac_opt of SOME tac => tac ctxt set_maps | NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt); val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_wit_thms set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); in map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] []) goals (map (fn tac => fn {context = ctxt, prems = _} => unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs) |> (fn thms => after_qed mk_wit_thms (map single thms) lthy) end) o prepare_def const_policy fact_policy internal qualify (K I) (K I) Ds map_b rel_b pred_b set_bs raw_csts; fun bnf_cmd (raw_csts, raw_plugins) = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) => let val plugins = raw_plugins lthy; val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_triv_wit_thms tac set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); val (mk_wit_thms, nontriv_wit_goals) = (case triv_tac_opt of NONE => (fn _ => [], map (map (rpair [])) wit_goalss) | SOME tac => (mk_triv_wit_thms tac, [])); in lthy |> Proof.theorem NONE (uncurry (register_bnf plugins key) oo after_qed mk_wit_thms) (map (single o rpair []) goals @ nontriv_wit_goals) |> Proof.unfolding ([[(@{thm OO_Grp_alt} :: @{thm mem_Collect_eq} :: defs, [])]]) |> Proof.refine_singleton (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (TRYALL (rtac ctxt refl)))) end) o prepare_def Do_Inline (user_policy Note_Some) false I Syntax.read_typ Syntax.read_term NONE Binding.empty Binding.empty Binding.empty [] raw_csts; fun print_bnfs ctxt = let fun pretty_set sets i = Pretty.block [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (nth sets i))]; fun pretty_bnf (key, BNF {T, map, sets, bd, live, lives, dead, deads, ...}) = Pretty.big_list (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)])) ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)], Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)], Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt map)]] @ List.map (pretty_set sets) (0 upto length sets - 1) @ [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt bd)]]); in Pretty.big_list "Registered bounded natural functors:" (map pretty_bnf (sort_by fst (Symtab.dest (Data.get (Context.Proof ctxt))))) |> Pretty.writeln end; val _ = Outer_Syntax.command \<^command_keyword>\print_bnfs\ "print all bounded natural functors" (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of))); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\bnf\ "register a type as a bounded natural functor" (parse_opt_binding_colon -- Parse.typ --| (Parse.reserved "map" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "sets" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) [] --| (Parse.reserved "bd" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "wits" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "rel" || Parse.reserved "plugins") Parse.term)) [] -- Scan.option ((Parse.reserved "rel" -- \<^keyword>\:\) |-- Parse.term) -- Scan.option ((Parse.reserved "pred" -- \<^keyword>\:\) |-- Parse.term) -- Scan.optional Plugin_Name.parse_filter (K Plugin_Name.default_filter) >> bnf_cmd); end; diff --git a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML @@ -1,2927 +1,2927 @@ (* Title: HOL/Tools/BNF/bnf_fp_def_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013, 2014 Sugared datatype and codatatype constructions. *) signature BNF_FP_DEF_SUGAR = sig type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list} type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list} type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list} type fp_sugar = {T: typ, BT: typ, X: typ, fp: BNF_Util.fp_kind, fp_res_index: int, fp_res: BNF_FP_Util.fp_result, pre_bnf: BNF_Def.bnf, fp_bnf: BNF_Def.bnf, absT_info: BNF_Comp.absT_info, fp_nesting_bnfs: BNF_Def.bnf list, live_nesting_bnfs: BNF_Def.bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option} val co_induct_of: 'a list -> 'a val strong_co_induct_of: 'a list -> 'a val morph_fp_bnf_sugar: morphism -> fp_bnf_sugar -> fp_bnf_sugar val morph_fp_co_induct_sugar: morphism -> fp_co_induct_sugar -> fp_co_induct_sugar val morph_fp_ctr_sugar: morphism -> fp_ctr_sugar -> fp_ctr_sugar val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar val fp_sugar_of: Proof.context -> string -> fp_sugar option val fp_sugar_of_global: theory -> string -> fp_sugar option val fp_sugars_of: Proof.context -> fp_sugar list val fp_sugars_of_global: theory -> fp_sugar list val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) -> theory -> theory val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val merge_type_args: BNF_Util.fp_kind -> ''a list * ''a list -> ''a list val type_args_named_constrained_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'a val type_binding_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'b val mixfix_of_spec: ((('a * 'b) * 'c) * 'd) * 'e -> 'b val mixfixed_ctr_specs_of_spec: (('a * 'b) * 'c) * 'd -> 'b val map_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'b val rel_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'c val pred_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'd val sel_default_eqs_of_spec: 'a * 'b -> 'b val mk_parametricity_goal: Proof.context -> term list -> term -> term -> term val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list -> 'a list val mk_ctor: typ list -> term -> term val mk_dtor: typ list -> term -> term val mk_bnf_sets: BNF_Def.bnf -> string * term list val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list val massage_simple_notes: string -> (bstring * 'a list * (int -> 'b)) list -> ((binding * 'c list) * ('a list * 'b) list) list val massage_multi_notes: string list -> typ list -> (string * 'a list list * (string -> 'b)) list -> ((binding * 'b) * ('a list * 'c list) list) list val define_ctrs_dtrs_for_type: string -> typ -> term -> term -> thm -> thm -> int -> int list -> term -> binding list -> mixfix list -> typ list list -> local_theory -> (term list list * term list * thm * thm list) * local_theory val wrap_ctrs: (string -> bool) -> BNF_Util.fp_kind -> bool -> string -> thm -> int -> int list -> thm -> thm -> binding list -> binding list list -> term list -> term list -> thm -> thm list -> local_theory -> Ctr_Sugar.ctr_sugar * local_theory val derive_map_set_rel_pred_thms: (string -> bool) -> BNF_Util.fp_kind -> int -> typ list -> typ list -> typ -> typ -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> string -> BNF_Def.bnf -> BNF_Def.bnf list -> typ -> term -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm -> thm list -> thm list -> thm list -> typ list list -> Ctr_Sugar.ctr_sugar -> local_theory -> (thm list * thm list * thm list list * thm list * thm list * thm list * thm list * thm list * thm list * thm list * thm list list list list * thm list list list list * thm list * thm list * thm list * thm list * thm list) * local_theory type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list) val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list) val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms val mk_co_recs_prelims: Proof.context -> BNF_Util.fp_kind -> typ list list list -> typ list -> typ list -> typ list -> typ list -> int list -> int list list -> term list -> term list * (typ list list * typ list list list list * term list list * term list list list list) option * (string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list)) option val repair_nullary_single_ctr: typ list list -> typ list list val mk_corec_p_pred_types: typ list -> int list -> typ list list val mk_corec_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list -> int list list -> term -> typ list list * (typ list list list list * typ list list list * typ list list list list * typ list) val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory -> (term * thm) * local_theory val define_rec: typ list list * typ list list list list * term list list * term list list list list -> (string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context -> (term * thm) * Proof.context val define_corec: 'a * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list -> term list -> term -> local_theory -> (term * thm) * local_theory val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list -> (string * term list) list -> term -> term -> typ list -> typ list -> term list * ((term * (term * term)) list * (int * term)) list * term val finish_induct_prem: Proof.context -> int -> term list -> term list * ((term * (term * term)) list * (int * term)) list * term -> term val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term -> term -> term -> int -> term list -> term list list -> term list -> term list list -> typ list list -> term val mk_induct_attrs: term list list -> Token.src list val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list -> Token.src list * Token.src list val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list -> ('a * typ list list list list * term list list * 'b) option -> thm -> thm list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list -> term list -> thm list -> Proof.context -> lfp_sugar_thms val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list -> thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> (thm list * thm) list val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list -> string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> int list list -> int list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list -> thm list -> gfp_sugar_thms val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> Ctr_Sugar.ctr_options * ((((((binding option * (typ * sort)) list * binding) * mixfix) * ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * term list) list -> local_theory -> local_theory val co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * Proof.context) -> ((Proof.context -> Plugin_Name.filter) * bool) * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list -> Proof.context -> local_theory val parse_ctr_arg: (binding * string) parser val parse_ctr_specs: ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list parser val parse_spec: ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) parser val parse_co_datatype: (Ctr_Sugar.ctr_options_cmd * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list) parser val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> (local_theory -> local_theory) parser end; structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR = struct open Ctr_Sugar open BNF_FP_Rec_Sugar_Util open BNF_Util open BNF_Comp open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar_Tactics val Eq_prefix = "Eq_"; val case_transferN = "case_transfer"; val ctor_iff_dtorN = "ctor_iff_dtor"; val ctr_transferN = "ctr_transfer"; val disc_transferN = "disc_transfer"; val sel_transferN = "sel_transfer"; val corec_codeN = "corec_code"; val corec_transferN = "corec_transfer"; val map_disc_iffN = "map_disc_iff"; val map_o_corecN = "map_o_corec"; val map_selN = "map_sel"; val pred_injectN = "pred_inject"; val rec_o_mapN = "rec_o_map"; val rec_transferN = "rec_transfer"; val set0N = "set0"; val set_casesN = "set_cases"; val set_introsN = "set_intros"; val set_inductN = "set_induct"; val set_selN = "set_sel"; type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list}; type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list}; type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list}; type fp_sugar = {T: typ, BT: typ, X: typ, fp: fp_kind, fp_res_index: int, fp_res: fp_result, pre_bnf: bnf, fp_bnf: bnf, absT_info: absT_info, fp_nesting_bnfs: bnf list, live_nesting_bnfs: bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option}; fun co_induct_of (i :: _) = i; fun strong_co_induct_of [_, s] = s; fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases} : fp_bnf_sugar) = {map_thms = map (Morphism.thm phi) map_thms, map_disc_iffs = map (Morphism.thm phi) map_disc_iffs, map_selss = map (map (Morphism.thm phi)) map_selss, rel_injects = map (Morphism.thm phi) rel_injects, rel_distincts = map (Morphism.thm phi) rel_distincts, rel_sels = map (Morphism.thm phi) rel_sels, rel_intros = map (Morphism.thm phi) rel_intros, rel_cases = map (Morphism.thm phi) rel_cases, pred_injects = map (Morphism.thm phi) pred_injects, set_thms = map (Morphism.thm phi) set_thms, set_selssss = map (map (map (map (Morphism.thm phi)))) set_selssss, set_introssss = map (map (map (map (Morphism.thm phi)))) set_introssss, set_cases = map (Morphism.thm phi) set_cases}; fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms, co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts} : fp_co_induct_sugar) = {co_rec = Morphism.term phi co_rec, common_co_inducts = map (Morphism.thm phi) common_co_inducts, co_inducts = map (Morphism.thm phi) co_inducts, co_rec_def = Morphism.thm phi co_rec_def, co_rec_thms = map (Morphism.thm phi) co_rec_thms, co_rec_discs = map (Morphism.thm phi) co_rec_discs, co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs, co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss, co_rec_codes = map (Morphism.thm phi) co_rec_codes, co_rec_transfers = map (Morphism.thm phi) co_rec_transfers, co_rec_o_maps = map (Morphism.thm phi) co_rec_o_maps, common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts, rel_co_inducts = map (Morphism.thm phi) rel_co_inducts, common_set_inducts = map (Morphism.thm phi) common_set_inducts, set_inducts = map (Morphism.thm phi) set_inducts}; fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctor_iff_dtor, ctr_defs, ctr_sugar, ctr_transfers, case_transfers, disc_transfers, sel_transfers} : fp_ctr_sugar) = {ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss, ctor_iff_dtor = Morphism.thm phi ctor_iff_dtor, ctr_defs = map (Morphism.thm phi) ctr_defs, ctr_sugar = morph_ctr_sugar phi ctr_sugar, ctr_transfers = map (Morphism.thm phi) ctr_transfers, case_transfers = map (Morphism.thm phi) case_transfers, disc_transfers = map (Morphism.thm phi) disc_transfers, sel_transfers = map (Morphism.thm phi) sel_transfers}; fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar, fp_co_induct_sugar} : fp_sugar) = {T = Morphism.typ phi T, BT = Morphism.typ phi BT, X = Morphism.typ phi X, fp = fp, fp_res = morph_fp_result phi fp_res, fp_res_index = fp_res_index, pre_bnf = morph_bnf phi pre_bnf, fp_bnf = morph_bnf phi fp_bnf, absT_info = morph_absT_info phi absT_info, fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs, live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs, fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar, fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar, fp_co_induct_sugar = Option.map (morph_fp_co_induct_sugar phi) fp_co_induct_sugar}; val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = fp_sugar Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun fp_sugar_of_generic context = Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); fun fp_sugars_of_generic context = Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; val fp_sugar_of = fp_sugar_of_generic o Context.Proof; val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory; val fp_sugars_of = fp_sugars_of_generic o Context.Proof; val fp_sugars_of_global = fp_sugars_of_generic o Context.Theory; structure FP_Sugar_Plugin = Plugin(type T = fp_sugar list); fun fp_sugars_interpretation name f = FP_Sugar_Plugin.interpretation name (fn fp_sugars => fn lthy => f (map (transfer_fp_sugar (Proof_Context.theory_of lthy)) fp_sugars) lthy); val interpret_fp_sugars = FP_Sugar_Plugin.data; val register_fp_sugars_raw = fold (fn fp_sugar as {T = Type (s, _), ...} => - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar)))); fun register_fp_sugars plugins fp_sugars = register_fp_sugars_raw fp_sugars #> interpret_fp_sugars plugins fp_sugars; fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs map_thmss common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss co_rec_disc_iffss co_rec_codess co_rec_transferss common_rel_co_inducts rel_co_inductss common_set_inducts set_inductss co_rec_o_mapss noted = let val fp_sugars = map_index (fn (kk, T) => {T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, fp_bnf = nth (#bnfs fp_res) kk, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = nth ctrXs_Tsss kk, ctor_iff_dtor = nth ctor_iff_dtors kk, ctr_defs = nth ctr_defss kk, ctr_sugar = nth ctr_sugars kk, ctr_transfers = nth ctr_transferss kk, case_transfers = nth case_transferss kk, disc_transfers = nth disc_transferss kk, sel_transfers = nth sel_transferss kk}, fp_bnf_sugar = {map_thms = nth map_thmss kk, map_disc_iffs = nth map_disc_iffss kk, map_selss = nth map_selsss kk, rel_injects = nth rel_injectss kk, rel_distincts = nth rel_distinctss kk, rel_sels = nth rel_selss kk, rel_intros = nth rel_intross kk, rel_cases = nth rel_casess kk, pred_injects = nth pred_injectss kk, set_thms = nth set_thmss kk, set_selssss = nth set_selsssss kk, set_introssss = nth set_introsssss kk, set_cases = nth set_casess kk}, fp_co_induct_sugar = SOME {co_rec = nth co_recs kk, common_co_inducts = common_co_inducts, co_inducts = nth co_inductss kk, co_rec_def = nth co_rec_defs kk, co_rec_thms = nth co_rec_thmss kk, co_rec_discs = nth co_rec_discss kk, co_rec_disc_iffs = nth co_rec_disc_iffss kk, co_rec_selss = nth co_rec_selsss kk, co_rec_codes = nth co_rec_codess kk, co_rec_transfers = nth co_rec_transferss kk, co_rec_o_maps = nth co_rec_o_mapss kk, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = nth rel_co_inductss kk, common_set_inducts = common_set_inducts, set_inducts = nth set_inductss kk}} |> morph_fp_sugar (substitute_noted_thm noted)) Ts; in register_fp_sugars_raw fp_sugars #> fold (interpret_bnf plugins) (#bnfs fp_res) #> interpret_fp_sugars plugins fp_sugars end; fun quasi_unambiguous_case_names names = let val ps = map (`Long_Name.base_name) names; val dups = Library.duplicates (op =) (map fst ps); fun underscore s = let val ss = Long_Name.explode s in space_implode "_" (drop (length ss - 2) ss) end; in map (fn (base, full) => if member (op =) dups base then underscore full else base) ps |> Name.variant_list [] end; fun zipper_map f = let fun zed _ [] = [] | zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys; in zed [] end; fun cannot_merge_types fp = error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters"); fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp; fun merge_type_args fp (As, As') = if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp; fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs; fun type_binding_of_spec (((((_, b), _), _), _), _) = b; fun mixfix_of_spec ((((_, mx), _), _), _) = mx; fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs; fun map_binding_of_spec ((_, (b, _, _)), _) = b; fun rel_binding_of_spec ((_, (_, b, _)), _) = b; fun pred_binding_of_spec ((_, (_, _, b)), _) = b; fun sel_default_eqs_of_spec (_, ts) = ts; fun ctr_sugar_kind_of_fp_kind Least_FP = Datatype | ctr_sugar_kind_of_fp_kind Greatest_FP = Codatatype; fun uncurry_thm 0 thm = thm | uncurry_thm 1 thm = thm | uncurry_thm n thm = rotate_prems ~1 (uncurry_thm (n - 1) (rotate_prems 1 (conjI RS thm))); fun choose_binary_fun fs AB = find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs; fun build_binary_fun_app fs t u = Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u)); fun build_the_rel ctxt Rs Ts A B = build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B); fun build_rel_app ctxt Rs Ts t u = build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u; fun build_set_app ctxt A t = Term.betapply (build_set ctxt A (fastype_of t), t); fun mk_parametricity_goal ctxt Rs t u = let val prem = build_the_rel ctxt Rs [] (fastype_of t) (fastype_of u) in HOLogic.mk_Trueprop (prem $ t $ u) end; val name_of_set = name_of_const "set function" domain_type; val fundefcong_attrs = @{attributes [fundef_cong]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs)); fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss); fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) = p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss; fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) = Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1)); fun flip_rels ctxt n thm = let val Rs = Term.add_vars (Thm.prop_of thm) []; val Rs' = rev (drop (length Rs - n) Rs); in infer_instantiate ctxt (map (fn f => (fst f, Thm.cterm_of ctxt (mk_flip f))) Rs') thm end; fun mk_ctor_or_dtor get_T Ts t = let val Type (_, Ts0) = get_T (fastype_of t) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_ctor = mk_ctor_or_dtor range_type; val mk_dtor = mk_ctor_or_dtor domain_type; fun mk_bnf_sets bnf = let val Type (T_name, Us) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val sets = sets_of_bnf bnf; fun mk_set U = (case find_index (curry (op =) U) lives of ~1 => Term.dummy | i => nth sets i); in (T_name, map mk_set Us) end; fun mk_xtor_co_recs thy fp fpTs Cs ts0 = let val nn = length fpTs; val (fpTs0, Cs0) = map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0 |> split_list; val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs); in map (Term.subst_TVars rho) ts0 end; fun liveness_of_fp_bnf n bnf = (case T_of_bnf bnf of Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts | _ => replicate n false); fun add_nesting_bnf_names Us = let fun add (Type (s, Ts)) ss = let val (needs, ss') = fold_map add Ts ss in if exists I needs then (true, insert (op =) s ss') else (false, ss') end | add T ss = (member (op =) Us T, ss); in snd oo add end; fun nesting_bnfs ctxt ctr_Tsss Us = map_filter (bnf_of ctxt) (fold (fold (fold (add_nesting_bnf_names Us))) ctr_Tsss []); fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p; fun massage_simple_notes base = filter_out (null o #2) #> map (fn (thmN, thms, f_attrs) => ((Binding.qualify true base (Binding.name thmN), []), map_index (fn (i, thm) => ([thm], f_attrs i)) thms)); fun massage_multi_notes b_names Ts = maps (fn (thmN, thmss, attrs) => @{map 3} (fn b_name => fn Type (T_name, _) => fn thms => ((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])])) b_names Ts thmss) #> filter_out (null o fst o hd o snd); fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy = let val ctor_absT = domain_type (fastype_of ctor); val (((w, xss), u'), _) = lthy |> yield_singleton (mk_Frees "w") ctor_absT ||>> mk_Freess "x" ctr_Tss ||>> yield_singleton Variable.variant_fixes fp_b_name; val u = Free (u', fpT); val ctor_iff_dtor_thm = let val goal = fold_rev Logic.all [w, u] (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT]) (Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor) |> Thm.close_derivation \<^here> end; val ctr_rhss = map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs)) ks xss; val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn b => fn mx => fn rhs => Local_Theory.define ((b, mx), ((Thm.make_def_binding (Config.get lthy bnf_internals) b, []), rhs)) #>> apsnd snd) ctr_bindings ctr_mixfixes ctr_rhss ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val ctr_defs = map (Morphism.thm phi) raw_ctr_defs; val ctrs0 = map (Morphism.term phi) raw_ctrs; in ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) end; fun wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs lthy = let val sumEN_thm' = unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms); fun exhaust_tac {context = ctxt, prems = _} = mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'; val inject_tacss = map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} => mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms; val half_distinct_tacss = map (map (fn (def, def') => fn {context = ctxt, ...} => mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def'])) (mk_half_pairss (`I ctr_defs)); val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss; fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs); val ctr_specs = @{map 3} ctr_spec_of disc_bindings ctrs0 sel_bindingss; val (ctr_sugar as {case_cong, ...}, lthy) = free_constructors (ctr_sugar_kind_of_fp_kind fp) tacss ((((plugins, discs_sels), standard_binding), ctr_specs), sel_default_eqs) lthy; val anonymous_notes = [([case_cong], fundefcong_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = if Config.get lthy bnf_internals then [(ctor_iff_dtorN, [ctor_iff_dtor_thm], K [])] |> massage_simple_notes fp_b_name else []; in (ctr_sugar, lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd) end; fun derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps fp_nested_rel_eq_onps fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm extra_unfolds_map extra_unfolds_set extra_unfolds_rel ctr_Tss ({casex, case_thms, discs, selss, sel_defs, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss, injects, distincts, distinct_discsss, ...} : ctr_sugar) lthy = let val n = length ctr_Tss; val ms = map length ctr_Tss; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpBT = B_ify_T fpT; val live_AsBs = filter (op <>) (As ~~ Bs); val live_As = map fst live_AsBs; val fTs = map (op -->) live_AsBs; val ((((((((xss, yss), fs), Ps), Rs), ta), tb), thesis), names_lthy) = lthy |> fold (fold Variable.declare_typ) [As, Bs] |> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss) ||>> mk_Frees "f" fTs ||>> mk_Frees "P" (map mk_pred1T live_As) ||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> yield_singleton (mk_Frees "a") fpT ||>> yield_singleton (mk_Frees "b") fpBT ||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT; val ctrAs = map (mk_ctr As) ctrs; val ctrBs = map (mk_ctr Bs) ctrs; val ctr_defs' = map2 (fn m => fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) ms ctr_defs; val ABfs = live_AsBs ~~ fs; fun derive_rel_case relAsBs rel_inject_thms rel_distinct_thms = let val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb; fun mk_assms ctrA ctrB ctxt = let val argA_Ts = binder_types (fastype_of ctrA); val argB_Ts = binder_types (fastype_of ctrB); val ((argAs, argBs), names_ctxt) = ctxt |> mk_Frees "x" argA_Ts ||>> mk_Frees "y" argB_Ts; val ctrA_args = list_comb (ctrA, argAs); val ctrB_args = list_comb (ctrB, argBs); in (fold_rev Logic.all (argAs @ argBs) (Logic.list_implies (mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) :: map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs, thesis)), names_ctxt) end; val (assms, names_lthy) = @{fold_map 2} mk_assms ctrAs ctrBs names_lthy; val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_case_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; fun derive_case_transfer rel_case_thm = let val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy; val caseA = mk_case As C casex; val caseB = mk_case Bs E casex; val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB; in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_case_transfer_tac ctxt rel_case_thm case_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; in if live = 0 then if plugins transfer_plugin then let val relAsBs = HOLogic.eq_const fpT; val rel_case_thm = derive_rel_case relAsBs [] []; val case_transfer_thm = derive_case_transfer rel_case_thm; val notes = [(case_transferN, [case_transfer_thm], K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> Local_Theory.notes notes; val subst = Morphism.thm (substitute_noted_thm noted); in (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [subst case_transfer_thm], [], []), lthy') end else (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy) else let val mapx = mk_map live As Bs (map_of_bnf fp_bnf); val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf); val setAs = map (mk_set As) (sets_of_bnf fp_bnf); val discAs = map (mk_disc_or_sel As) discs; val discBs = map (mk_disc_or_sel Bs) discs; val selAss = map (map (mk_disc_or_sel As)) selss; val selBss = map (map (mk_disc_or_sel Bs)) selss; val map_ctor_thm = if fp = Least_FP then fp_map_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val (y as Free (y_s, _), _) = lthy |> yield_singleton (mk_Frees "y") y_T; val ctor_cong = infer_instantiate' lthy [NONE, NONE, SOME (Thm.cterm_of lthy ctorB)] arg_cong; val fp_map_thm' = fp_map_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y))]) |> unfold_thms lthy [dtor_ctor]; in (fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans)) |> Drule.generalize (Names.empty, Names.make1_set y_s) end; val map_thms = let fun mk_goal ctrA ctrB xs ys = let val fmap = list_comb (mapx, fs); fun mk_arg (x as Free (_, T)) (Free (_, U)) = if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x; val xs' = map2 mk_arg xs ys; in mk_Trueprop_eq (fmap $ list_comb (ctrA, xs), list_comb (ctrB, xs')) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_tac ctxt abs_inverses pre_map_def map_ctor_thm live_nesting_map_id0s ctr_defs' extra_unfolds_map) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val set0_thms = let fun mk_goal A setA ctrA xs = let val sets = map (build_set_app lthy A) (filter (exists_subtype_in [A] o fastype_of) xs); in mk_Trueprop_eq (setA $ list_comb (ctrA, xs), (if null sets then HOLogic.mk_set A [] else Library.foldl1 mk_union sets)) end; val goals = @{map 2} (fn live_A => fn setA => map2 (mk_goal live_A setA) ctrAs xss) live_As setAs |> flat; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_set_thms fp_nesting_set_maps live_nesting_set_maps ctr_defs' extra_unfolds_set) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_thms = set0_thms |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left}); val rel_ctor_thm = if fp = Least_FP then fp_rel_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val z_T = domain_type (fastype_of ctorB); val ((y as Free (y_s, _), z as Free (z_s, _)), _) = lthy |> yield_singleton (mk_Frees "y") y_T ||>> yield_singleton (mk_Frees "z") z_T; in fp_rel_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))]) |> unfold_thms lthy [dtor_ctor] |> Drule.generalize (Names.empty, Names.make2_set y_s z_s) end; val rel_inject_thms = let fun mk_goal ctrA ctrB xs ys = let val lhs = list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys); val conjuncts = map2 (build_rel_app lthy Rs []) xs ys; in HOLogic.mk_Trueprop (if null conjuncts then lhs else HOLogic.mk_eq (lhs, Library.foldr1 HOLogic.mk_conj conjuncts)) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val half_rel_distinct_thmss = let fun mk_goal ((ctrA, xs), (ctrB, ys)) = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys))); val rel_infos = (ctrAs ~~ xss, ctrBs ~~ yss); val goalss = map (map mk_goal) (mk_half_pairss rel_infos); val goals = flat goalss; in unflat goalss (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_flip = rel_flip_of_bnf fp_bnf; fun mk_other_half_rel_distinct_thm thm = flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2); val other_half_rel_distinct_thmss = map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss; val (rel_distinct_thms, _) = join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss; fun mk_rel_intro_thm m thm = uncurry_thm m (thm RS iffD2) handle THM _ => thm; val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms; val rel_code_thms = map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @ map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms; val ctr_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val (set_cases_thms, set_cases_attrss) = let fun mk_prems assms elem t ctxt = (case fastype_of t of Type (type_name, xs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val new_var = not (T = fastype_of elem); val (x, ctxt') = if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt); in mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt' |>> map (new_var ? Logic.all x) end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt)) | T => rpair ctxt (if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else [])); in split_list (map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)); val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy; val premss = map (fn ctr => let val (args, names_lthy) = mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy; in flat (zipper_map (fn (prev_args, arg, next_args) => let val (args_with_elem, args_without_elem) = if fastype_of arg = A then (prev_args @ [elem] @ next_args, prev_args @ next_args) else `I (prev_args @ [arg] @ next_args); in mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))] elem arg names_lthy |> fst |> map (fold_rev Logic.all args_without_elem) end) args) end) ctrAs; val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars (flat premss) goal (fn {context = ctxt, prems} => mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms) |> Thm.close_derivation \<^here> |> rotate_prems ~1; val cases_set_attr = - Attrib.internal (K (Induct.cases_pred (name_of_set set))); + Attrib.internal \<^here> (K (Induct.cases_pred (name_of_set set))); val ctr_names = quasi_unambiguous_case_names (flat (map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs))); in (* TODO: @{attributes [elim?]} *) (thm, [Attrib.consumes 1, cases_set_attr, Attrib.case_names ctr_names]) end) setAs) end; val (set_intros_thmssss, set_intros_thms) = let fun mk_goals A setA ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (y, ctxt') = yield_singleton (mk_Frees "y") T ctxt; val assm = mk_Trueprop_mem (y, set $ t); in apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args y ctxt') end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt)) | T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt)); val (goalssss, _) = fold_map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)) in @{fold_map 2} (fn ctr => fn xs => fold_map (mk_goals A set (Term.list_comb (ctr, xs))) xs) ctrAs xss end) setAs lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_sel_thms = let val n = length discAs; fun mk_conjunct n k discA selAs discB selBs = (if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @ (if null selAs then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA $ ta, discB $ tb], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs []) (map (rapp ta) selAs) (map (rapp tb) selBs)))]); val goals = if n = 0 then [] else [mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb, (case flat (@{map 5} (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of [] => \<^term>\True\ | conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))]; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs)) |> Thm.close_derivation \<^here>; in map prove goals end; val (rel_case_thm, rel_case_attrs) = let val thm = derive_rel_case relAsBs rel_inject_thms rel_distinct_thms; val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs); in (thm, [Attrib.case_names ctr_names, Attrib.consumes 1] @ @{attributes [cases pred]}) end; val case_transfer_thm = derive_case_transfer rel_case_thm; val sel_transfer_thms = if null selAss then [] else let val shared_sels = foldl1 (uncurry (inter (op =))) (map (op ~~) (selAss ~~ selBss)); val goals = map (uncurry (mk_parametricity_goal names_lthy Rs)) shared_sels; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sel_transfer_tac ctxt n sel_defs case_transfer_thm) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val disc_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt (the_single rel_sel_thms) (the_single exhaust_discs) (flat (flat distinct_discsss))) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val map_disc_iff_thms = let val discsB = map (mk_disc_or_sel Bs) discs; val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs; fun mk_goal (discA_t, discB) = if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then NONE else SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t)); val goals = map_filter mk_goal (discsA_t ~~ discsB); in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val (map_sel_thmss, map_sel_thms) = let fun mk_goal discA selA selB = let val prem = Term.betapply (discA, ta); val lhs = selB $ (Term.list_comb (mapx, fs) $ ta); val lhsT = fastype_of lhs; val map_rhsT = map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT; val map_rhs = build_map lthy [] [] (the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT); val rhs = (case map_rhs of Const (\<^const_name>\id\, _) => selA $ ta | _ => map_rhs $ (selA $ ta)); val concl = mk_Trueprop_eq (lhs, rhs); in if is_refl_bool prem then concl else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl) end; val goalss = @{map 3} (map2 o mk_goal) discAs selAss selBss; val goals = flat goalss; in `(unflat goalss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms (flat sel_thmss) live_nesting_map_id0s) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val (set_sel_thmssss, set_sel_thms) = let fun mk_goal setA discA selA ctxt = let val prem = Term.betapply (discA, ta); val sel_rangeT = range_type (fastype_of selA); val A = HOLogic.dest_setT (range_type (fastype_of setA)); fun travese_nested_types t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; val assm = mk_Trueprop_mem (x, set $ a); in travese_nested_types x ctxt' |>> map (Logic.mk_implies o pair assm) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt)); val (concls, ctxt') = if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt) else travese_nested_types (selA $ ta) ctxt; in if exists_subtype_in [A] sel_rangeT then if is_refl_bool prem then (concls, ctxt') else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt') else ([], ctxt) end; val (goalssss, _) = fold_map (fn set => @{fold_map 2} (fold_map o mk_goal set) discAs selAss) setAs names_lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val pred_injects = let val rel_eq_onp_with_tops_of = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} lthy))); val eq_onps = map rel_eq_onp_with_tops_of (map rel_eq_onp_of_bnf fp_bnfs @ fp_nesting_rel_eq_onps @ live_nesting_rel_eq_onps @ fp_nested_rel_eq_onps); val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) live_As); val cts = map (SOME o Thm.cterm_of lthy) (map mk_eq_onp Ps); val get_rhs = Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> snd; val pred_eq_onp_conj = List.foldr (fn (_, thm) => thm RS @{thm eq_onp_live_step}) @{thm refl[of True]}; fun predify_rel_inject rel_inject = let val conjuncts = try (get_rhs #> HOLogic.dest_conj) rel_inject |> the_default []; fun postproc thm = if null conjuncts then thm RS (@{thm eq_onp_same_args} RS iffD1) else @{thm box_equals} OF [thm, @{thm eq_onp_same_args}, pred_eq_onp_conj conjuncts |> unfold_thms lthy @{thms simp_thms(21)}]; in rel_inject |> Thm.instantiate' cTs cts |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))) |> unfold_thms lthy eq_onps |> postproc |> unfold_thms lthy @{thms top_conj} end; in rel_inject_thms |> map (unfold_thms lthy [@{thm conj_assoc}]) |> map predify_rel_inject |> Proof_Context.export names_lthy lthy end; val anonymous_notes = [(rel_code_thms, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = (if Config.get lthy bnf_internals then [(set0N, set0_thms, K [])] else []) @ [(case_transferN, [case_transfer_thm], K []), (ctr_transferN, ctr_transfer_thms, K []), (disc_transferN, disc_transfer_thms, K []), (sel_transferN, sel_transfer_thms, K []), (mapN, map_thms, K (nitpicksimp_attrs @ simp_attrs)), (map_disc_iffN, map_disc_iff_thms, K simp_attrs), (map_selN, map_sel_thms, K []), (pred_injectN, pred_injects, K simp_attrs), (rel_casesN, [rel_case_thm], K rel_case_attrs), (rel_distinctN, rel_distinct_thms, K simp_attrs), (rel_injectN, rel_inject_thms, K simp_attrs), (rel_introsN, rel_intro_thms, K []), (rel_selN, rel_sel_thms, K []), (setN, set_thms, K (case_fp fp nitpicksimp_attrs [] @ simp_attrs)), (set_casesN, set_cases_thms, nth set_cases_attrss), (set_introsN, set_intros_thms, K []), (set_selN, set_sel_thms, K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) map_thms) |> fp = Least_FP ? uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) rel_code_thms) |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) set0_thms) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (rel_code_thms @ map_thms @ set_thms)) |> Local_Theory.notes (anonymous_notes @ notes); val subst = Morphism.thm (substitute_noted_thm noted); in ((map subst map_thms, map subst map_disc_iff_thms, map (map subst) map_sel_thmss, map subst rel_inject_thms, map subst rel_distinct_thms, map subst rel_sel_thms, map subst rel_intro_thms, [subst rel_case_thm], map subst pred_injects, map subst set_thms, map (map (map (map subst))) set_sel_thmssss, map (map (map (map subst))) set_intros_thmssss, map subst set_cases_thms, map subst ctr_transfer_thms, [subst case_transfer_thm], map subst disc_transfer_thms, map subst sel_transfer_thms), lthy') end end; type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list); fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) = ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs), (map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms; val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism; type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list); fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair), corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs), (corec_selsss, corec_sel_attrs)) = ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs, coinduct_attrs_pair), map (map (Morphism.thm phi)) corecss, map (map (Morphism.thm phi)) corec_discss, (map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs), (map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms; val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism; fun unzip_recT (Type (\<^type_name>\prod\, [_, TFree x])) (T as Type (\<^type_name>\prod\, Ts as [_, TFree y])) = if x = y then [T] else Ts | unzip_recT _ (Type (\<^type_name>\prod\, Ts as [_, TFree _])) = Ts | unzip_recT _ T = [T]; fun mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts = let val Css = map2 replicate ns Cs; val x_Tssss = @{map 6} (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_rec_fun_T => map2 (map2 unzip_recT) ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T))) absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts; val x_Tsss' = map (map flat_rec_arg_args) x_Tssss; val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css; val ((fss, xssss), _) = ctxt |> mk_Freess "f" f_Tss ||>> mk_Freessss "x" x_Tssss; in (f_Tss, x_Tssss, fss, xssss) end; fun unzip_corecT (Type (\<^type_name>\sum\, _)) T = [T] | unzip_corecT _ (Type (\<^type_name>\sum\, Ts)) = Ts | unzip_corecT _ T = [T]; (*avoid "'a itself" arguments in corecursors*) fun repair_nullary_single_ctr [[]] = [[HOLogic.unitT]] | repair_nullary_single_ctr Tss = Tss; fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts = let val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; val g_absTs = map range_type fun_Ts; val g_Tsss = map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) Cs ctr_Tsss' g_Tsss; val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; in (q_Tssss, g_Tsss, g_Tssss, g_absTs) end; fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs; fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec = (mk_corec_p_pred_types Cs ns, mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss (binder_fun_types (fastype_of dtor_corec))); fun mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts = let val p_Tss = mk_corec_p_pred_types Cs ns; val (q_Tssss, g_Tsss, g_Tssss, corec_types) = mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts; val (((((Free (x, _), cs), pss), qssss), gssss), _) = ctxt |> yield_singleton (mk_Frees "x") dummyT ||>> mk_Frees "a" Cs ||>> mk_Freess "p" p_Tss ||>> mk_Freessss "q" q_Tssss ||>> mk_Freessss "g" g_Tssss; val cpss = map2 (map o rapp) cs pss; fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd); fun build_dtor_corec_arg _ [] [cg] = cg | build_dtor_corec_arg T [cq] [cg, cg'] = mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg) (build_sum_inj Inr_const (fastype_of cg', T) $ cg'); val pgss = @{map 3} flat_corec_preds_predsss_gettersss pss qssss gssss; val cqssss = map2 (map o map o map o rapp) cs qssss; val cgssss = map2 (map o map o map o rapp) cs gssss; val cqgsss = @{map 3} (@{map 3} (@{map 3} build_dtor_corec_arg)) g_Tsss cqssss cgssss; in (x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types)) end; fun mk_co_recs_prelims ctxt fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 = let val thy = Proof_Context.theory_of ctxt; val (xtor_co_rec_fun_Ts, xtor_co_recs) = mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd); val (recs_args_types, corecs_args_types) = if fp = Least_FP then mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (rpair NONE o SOME) else mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (pair NONE o SOME); in (xtor_co_recs, recs_args_types, corecs_args_types) end; fun mk_preds_getterss_join c cps absT abs cqgss = let val n = length cqgss; val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqgss; in Term.lambda c (mk_IfN absT cps ts) end; fun define_co_rec_as fp Cs fpT b rhs lthy0 = let val thy = Proof_Context.theory_of lthy0; val ((cst, (_, def)), (lthy', lthy)) = lthy0 |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((b, NoSyn), ((Thm.make_def_binding (Config.get lthy0 bnf_internals) b, []), rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst); val def' = Morphism.thm phi def; in ((cst', def'), lthy') end; fun define_rec (_, _, fss, xssss) mk_binding fpTs Cs reps ctor_rec = let val nn = length fpTs; val (ctor_rec_absTs, fpT) = strip_typeN nn (fastype_of ctor_rec) |>> map domain_type ||> domain_type; in define_co_rec_as Least_FP Cs fpT (mk_binding recN) (fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_rec, @{map 4} (fn ctor_rec_absT => fn rep => fn fs => fn xsss => mk_case_absumprod ctor_rec_absT rep fs (map (map HOLogic.mk_tuple) xsss) (map flat_rec_arg_args xsss)) ctor_rec_absTs reps fss xssss))) end; fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = let val nn = length fpTs; val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); in define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN) (fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec, @{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss))) end; fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) = (case AList.lookup (op =) setss_fp_nesting T_name of NONE => [] | SOME raw_sets0 => let val (Xs_Ts, (Ts, raw_sets)) = filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0)) |> split_list ||> split_list; val sets = map (mk_set Ts0) raw_sets; val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts; val xysets = map (pair x) (ys ~~ sets); val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts; in flat (map2 (map o apfst o cons) xysets ppremss) end) | mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X = [([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))] | mk_induct_raw_prem_prems _ _ _ _ _ = []; fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts = let val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts; val pprems = flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts); val y = Term.list_comb (ctr, map alter_x xs); val p' = enforce_type names_ctxt domain_type (fastype_of y) p; in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end; fun close_induct_prem_prem nn ps xs t = fold_rev Logic.all (map Free (drop (nn + length xs) (rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t; fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) = let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) => mk_Trueprop_mem (y, set $ x')) xysets, HOLogic.mk_Trueprop (p' $ x))) end; fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) = fold_rev Logic.all xs (Logic.list_implies (map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl)); fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts = let fun build_the_rel T Xs_T = build_rel [] ctxt [] [] (fn (T, X) => nth rs' (find_index (fn Xs => member (op =) Xs X) Xss) |> enforce_type ctxt domain_type T) (T, Xs_T) |> Term.subst_atomic_types (flat Xss ~~ flat fpTss); fun build_rel_app usel vsel Xs_T = fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T); in (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @ (if null usels then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))]) end; fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss = @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n) (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss |> flat |> Library.foldr1 HOLogic.mk_conj handle List.Empty => \<^term>\True\; fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss = fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr, HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss))); fun postproc_co_induct ctxt nn prop prop_conj = Drule.zero_var_indexes #> `(conj_dests nn) #>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop)) ##> (fn thm => Thm.permute_prems 0 (~ nn) (if nn = 1 then thm RS prop else funpow nn (fn thm => unfold_thms ctxt @{thms conj_assoc} (thm RS prop_conj)) thm)); fun mk_induct_attrs ctrss = let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss); in [Attrib.case_names induct_cases] end; fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val fpB_Ts = map B_ify_T fpA_Ts; val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss; val ctrBss = map (map B_ify) ctrAss; val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> mk_Freesss "a" ctrAs_Tsss ||>> mk_Freesss "b" ctrBs_Tsss; val prems = let fun mk_prem ctrA ctrB argAs argBs = fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies) (map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs) (HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts (Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs))))); in flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs)); val vars = Variable.add_free_names ctxt goal []; val rel_induct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm, mk_induct_attrs ctrAss) end; fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions abs_inverses ctrss ctr_defss recs rec_defs ctxt = let val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss; val nn = length pre_bnfs; val ns = map length ctr_Tsss; val mss = map (map length) ctr_Tsss; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val (((ps, xsss), us'), names_ctxt) = ctxt |> mk_Frees "P" (map mk_pred1T fpTs) ||>> mk_Freesss "x" ctr_Tsss ||>> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs; val (induct_thms, induct_thm) = let val raw_premss = @{map 4} (@{map 3} o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting) ps ctrss ctr_Tsss ctrXs_Tsss; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)); val goal = Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps))) (raw_premss, concl); val vars = Variable.add_free_names ctxt goal []; val kksss = map (map (map (fst o snd) o #2)) raw_premss; val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss); val thm = Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses abs_inverses fp_nesting_set_maps pre_set_defss) |> Thm.close_derivation \<^here>; in `(conj_dests nn) thm end; val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss; fun mk_rec_thmss (_, x_Tssss, fss, _) recs rec_defs ctor_rec_thms = let val frecs = map (lists_bmoc fss) recs; fun mk_goal frec xctr f xs fxs = fold_rev (fold_rev Logic.all) (xs :: fss) (mk_Trueprop_eq (frec $ xctr, Term.list_comb (f, fxs))); fun maybe_tick (T, U) u f = if try (fst o HOLogic.dest_prodT) U = SOME T then Term.lambda u (HOLogic.mk_prod (u, f $ u)) else f; fun build_rec (x as Free (_, T)) U = if T = U then x else let val build_simple = indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs (fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk)); in build_map ctxt [] [] build_simple (T, U) $ x end; val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss; val goalss = @{map 5} (@{map 4} o mk_goal) frecs xctrss fss xsss fxsss; val tacss = @{map 4} (map ooo mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs) ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss end; val rec_thmss = mk_rec_thmss (the rec_args_typess) recs rec_defs ctor_rec_thms; in ((induct_thms, induct_thm, mk_induct_attrs ctrss), (rec_thmss, nitpicksimp_attrs @ simp_attrs)) end; fun mk_coinduct_attrs fpTs ctrss discss mss = let val fp_b_names = map base_name_of_typ fpTs; fun mk_coinduct_concls ms discs ctrs = let fun mk_disc_concl disc = [name_of_disc disc]; fun mk_ctr_concl 0 _ = [] | mk_ctr_concl _ ctr = [name_of_ctr ctr]; val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]]; val ctr_concls = map2 mk_ctr_concl ms ctrs; in flat (map2 append disc_concls ctr_concls) end; val coinduct_cases = quasi_unambiguous_case_names (map (prefix Eq_prefix) fp_b_names); val coinduct_conclss = @{map 3} (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss; val coinduct_case_names_attr = Attrib.case_names coinduct_cases; val coinduct_case_concl_attrs = map2 (fn casex => fn concls => Attrib.case_conclusion (casex, concls)) coinduct_cases coinduct_conclss; val common_coinduct_attrs = coinduct_case_names_attr :: coinduct_case_concl_attrs; val coinduct_attrs = Attrib.consumes 1 :: coinduct_case_names_attr :: coinduct_case_concl_attrs; in (coinduct_attrs, common_coinduct_attrs) end; fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list) abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpB_Ts = map B_ify_T fpA_Ts; val (Rs, IRs, fpAs, fpBs, _) = let val fp_names = map base_name_of_typ fpA_Ts; val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> Variable.variant_fixes fp_names ||>> Variable.variant_fixes (map (suffix "'") fp_names); in (Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts, names_ctxt) end; val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) = let val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss); fun mk_selsss ts Ts = map2 (map o map o rapp) ts (map (map (map (mk_disc_or_sel Ts))) selsss); in ((mk_discss fpAs As, mk_selsss fpAs As), (mk_discss fpBs Bs, mk_selsss fpBs Bs)) end; val prems = let fun mk_prem_ctr_concls n k discA_t selA_ts discB_t selB_ts = (if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @ (case (selA_ts, selB_ts) of ([], []) => [] | (_ :: _, _ :: _) => [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA_t, discB_t], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]); fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss = Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n) (1 upto n) discA_ts selA_tss discB_ts selB_tss)) handle List.Empty => \<^term>\True\; fun mk_prem IR tA tB n discA_ts selA_tss discB_ts selB_tss = fold_rev Logic.all [tA, tB] (Logic.mk_implies (HOLogic.mk_Trueprop (IR $ tA $ tB), HOLogic.mk_Trueprop (mk_prem_concl n discA_ts selA_tss discB_ts selB_tss))); in @{map 8} mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts))); val vars = Variable.add_free_names ctxt goal []; val rel_coinduct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars) (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm, mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss) end; fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = let fun mk_prems A Ps ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val X = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt; val assm = mk_Trueprop_mem (x, set $ a); in (case build_binary_fun_app Ps x a of NONE => mk_prems A Ps ctr_args x ctxt' |>> map (Logic.all x o Logic.mk_implies o pair assm) | SOME f => ([Logic.all x (Logic.mk_implies (assm, Logic.mk_implies (HOLogic.mk_Trueprop f, HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))], ctxt')) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt) else ([], ctxt)); fun mk_prems_for_ctr A Ps ctr ctxt = let val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt; in fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt' |>> map (fold_rev Logic.all args) o flat |>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr))) end; fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt = let val ((x, fp), ctxt') = ctxt |> yield_singleton (mk_Frees "x") A ||>> yield_singleton (mk_Frees "a") fpT; val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x) (the (build_binary_fun_app Ps x fp))); in fold_map (mk_prems_for_ctr A Ps) ctrs ctxt' |>> split_list |>> map_prod flat flat |>> apfst (rpair concl) end; fun mk_thm ctxt fpTs ctrss sets = let val A = HOLogic.dest_setT (range_type (fastype_of (hd sets))); val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt; val (((prems, concl), case_names), ctxt'') = fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt' |>> apfst split_list o split_list |>> apfst (apfst flat) |>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj)) |>> apsnd flat; val vars = fold (Variable.add_free_names ctxt) (concl :: prems) []; val thm = Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> Thm.close_derivation \<^here>; val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names); - val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets; + val induct_set_attrs = map (Attrib.internal \<^here> o K o Induct.induct_pred o name_of_set) sets; in (thm, case_names_attr :: induct_set_attrs) end val consumes_attr = Attrib.consumes 1; in map (mk_thm ctxt fpTs ctrss #> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr)) (transpose setss) end; fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt = let val n = Thm.nprems_of coind; val m = Thm.nprems_of (hd rel_monos) - n; fun mk_inst phi = (phi, Thm.cterm_of ctxt (mk_union (Var phi, HOLogic.eq_const (fst (dest_pred2T (#2 phi)))))); val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map mk_inst; fun mk_unfold rel_eq rel_mono = let val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl]; val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset}); in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS eqTrueI end; val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)}; in Thm.instantiate (TVars.empty, Vars.make insts) coind |> unfold_thms ctxt unfolds end; fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) = let val nn = length pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val exhausts = map #exhaust ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val (((rs, us'), vs'), _) = ctxt |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs) ||>> Variable.variant_fixes fp_b_names ||>> Variable.variant_fixes (map (suffix "'") fp_b_names); val us = map2 (curry Free) us' fpTs; val udiscss = map2 (map o rapp) us discss; val uselsss = map2 (map o map o rapp) us selsss; val vs = map2 (curry Free) vs' fpTs; val vdiscss = map2 (map o rapp) vs discss; val vselsss = map2 (map o map o rapp) vs selsss; val uvrs = @{map 3} (fn r => fn u => fn v => r $ u $ v) rs us vs; val uv_eqs = map2 (curry HOLogic.mk_eq) us vs; val strong_rs = @{map 4} (fn u => fn v => fn uvr => fn uv_eq => fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v))) uvrs us vs)) fun mk_goal rs0' = Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs) (map alter_r rs0')) uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss, concl); val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else [])); fun prove dtor_coinduct' goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss)) |> Thm.close_derivation \<^here>; val rel_eqs = map rel_eq_of_bnf pre_bnfs; val rel_monos = map rel_mono_of_bnf pre_bnfs; val dtor_coinducts = [dtor_coinduct] @ (if strong then [mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p ctxt] else []); in map2 (postproc_co_induct ctxt nn mp @{thm conj_commute[THEN iffD1]} oo prove) dtor_coinducts goals end; fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs = let fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec = iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]]; val ctor_dtor_corec_thms = @{map 3} mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms; val pre_map_defs = map map_def_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val ctrss = map #ctrs ctr_sugars; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val discIss = map #discIs ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss ctr_sugars; fun mk_maybe_not pos = not pos ? HOLogic.mk_not; val gcorecs = map (lists_bmoc pgss) corecs; val corec_thmss = let val (us', _) = ctxt |> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; fun mk_goal c cps gcorec n k ctr m cfs' = fold_rev (fold_rev Logic.all) ([c] :: pgss) (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps, mk_Trueprop_eq (gcorec $ c, Term.list_comb (ctr, take m cfs')))); val mk_U = typ_subst_nonatomic (map2 (fn C => fn fpT => (mk_sumT (fpT, C), fpT)) Cs fpTs); fun tack (c, u) f = let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in Term.lambda x' (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ x') end; fun build_corec cqg = let val T = fastype_of cqg in if exists_subtype_in Cs T then let val U = mk_U T; val build_simple = indexify fst (map2 (curry mk_sumT) fpTs Cs) (fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk)); in build_map ctxt [] [] build_simple (T, U) $ cqg end else cqg end; val cqgsss' = map (map (map build_corec)) cqgsss; val goalss = @{map 8} (@{map 4} oooo mk_goal) cs cpss gcorecs ns kss ctrss mss cqgsss'; val tacss = @{map 4} (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s) ctor_dtor_corec_thms pre_map_defs abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss |> map (map (unfold_thms ctxt @{thms case_sum_if})) end; val corec_disc_iff_thmss = let fun mk_goal c cps gcorec n k disc = mk_Trueprop_eq (disc $ (gcorec $ c), if n = 1 then \<^Const>\True\ else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps)); val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; fun mk_case_split' cp = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt cp)] @{thm case_split}; val case_splitss' = map (map mk_case_split') cpss; val tacss = @{map 3} (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss; fun prove goal tac = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (tac o #context)) |> Thm.close_derivation \<^here>; fun proves [_] [_] = [] | proves goals tacs = map2 prove goals tacs; in map2 proves goalss tacss end; fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs); val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss; fun mk_corec_sel_thm corec_thm sel sel_thm = let val (domT, ranT) = dest_funT (fastype_of sel); val arg_cong' = Thm.instantiate' (map (SOME o Thm.ctyp_of ctxt) [domT, ranT]) [NONE, NONE, SOME (Thm.cterm_of ctxt sel)] arg_cong |> Thm.varifyT_global; val sel_thm' = sel_thm RSN (2, trans); in corec_thm RS arg_cong' RS sel_thm' end; fun mk_corec_sel_thms corec_thmss = @{map 3} (@{map 3} (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss; val corec_sel_thmsss = mk_corec_sel_thms corec_thmss; in ((coinduct_thms_pairs, mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, simp_attrs), (corec_sel_thmsss, simp_attrs)) end; fun define_co_datatypes prepare_plugins prepare_constraint prepare_typ prepare_term fp construct_fp ((raw_plugins, discs_sels0), specs) lthy = let val plugins = prepare_plugins lthy raw_plugins; val discs_sels = discs_sels0 orelse fp = Greatest_FP; val nn = length specs; val fp_bs = map type_binding_of_spec specs; val fp_b_names = map Binding.name_of fp_bs; val fp_common_name = mk_common_name fp_b_names; val map_bs = map map_binding_of_spec specs; val rel_bs = map rel_binding_of_spec specs; val pred_bs = map pred_binding_of_spec specs; fun prepare_type_arg (_, (ty, c)) = let val TFree (s, _) = prepare_typ lthy ty in TFree (s, prepare_constraint lthy c) end; val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs; val unsorted_Ass0 = map (map (resort_tfree_or_tvar \<^sort>\type\)) Ass0; val unsorted_As = Library.foldr1 (merge_type_args fp) unsorted_Ass0; val num_As = length unsorted_As; val set_boss = map (map fst o type_args_named_constrained_of_spec) specs; val set_bss = map (map (the_default Binding.empty)) set_boss; fun add_fake_type spec = Typedecl.basic_typedecl {final = true} (type_binding_of_spec spec, num_As, Mixfix.reset_pos (mixfix_of_spec spec)); val (fake_T_names, fake_lthy) = fold_map add_fake_type specs lthy; val qsoty = quote o Syntax.string_of_typ fake_lthy; val _ = (case Library.duplicates (op =) unsorted_As of [] => () | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^ "datatype specification")); val bad_args = map (Logic.type_map (singleton (Variable.polymorphic lthy))) unsorted_As |> filter_out Term.is_TVar; val _ = null bad_args orelse error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^ "datatype specification"); val mixfixes = map mixfix_of_spec specs; val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => () | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b))); val mx_ctr_specss = map mixfixed_ctr_specs_of_spec specs; val ctr_specss = map (map fst) mx_ctr_specss; val ctr_mixfixess = map (map snd) mx_ctr_specss; val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss; val ctr_bindingss = map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of_ctr_spec)) fp_b_names ctr_specss; val ctr_argsss = map (map args_of_ctr_spec) ctr_specss; val sel_bindingsss = map (map (map fst)) ctr_argsss; val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss; val raw_sel_default_eqss = map sel_default_eqs_of_spec specs; val (As :: _) :: fake_ctr_Tsss = burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0); val As' = map dest_TFree As; val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss []; val _ = (case subtract (op =) As' rhs_As' of [] => () | extras => error ("Extra type variables on right-hand side: " ^ commas (map (qsoty o TFree) extras))); val fake_Ts = map (fn s => Type (s, As)) fake_T_names; val ((((Bs0, Cs), Es), Xs), _) = lthy |> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As |> mk_TFrees num_As ||>> mk_TFrees nn ||>> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) = s = s' andalso (Ts = Ts' orelse error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^ " (expected " ^ qsoty T' ^ ")")) | eq_fpT_check _ _ = false; fun freeze_fp (T as Type (s, Ts)) = (case find_index (eq_fpT_check T) fake_Ts of ~1 => Type (s, map freeze_fp Ts) | kk => nth Xs kk) | freeze_fp T = T; val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts); val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val _ = let fun add_deps i = fold (fn T => fold_index (fn (j, X) => (i <> j andalso exists_subtype_in [X] T) ? insert (op =) (i, j)) Xs); val add_missing_nodes = fold (AList.default (op =) o rpair []) (0 upto nn - 1); val deps = fold_index (uncurry (fold o add_deps)) ctrXs_Tsss [] |> AList.group (op =) |> add_missing_nodes; val G = Int_Graph.make (map (apfst (rpair ())) deps); val sccs = map (sort int_ord) (Int_Graph.strong_conn G); val str_of_scc = prefix (co_prefix fp ^ "datatype ") o space_implode " and " o map (suffix " = \" o Long_Name.base_name); fun warn [_] = () | warn sccs = warning ("Defined types not fully mutually " ^ co_prefix fp ^ "recursive\n\ \Alternative specification:\n" ^ cat_lines (map (prefix " " o str_of_scc o map (nth fp_b_names)) sccs)); in warn (order_strong_conn (op =) Int_Graph.make Int_Graph.topological_order deps sccs) end; val killed_As = map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE) (As ~~ transpose set_boss); val (((pre_bnfs, absT_infos), _), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels, xtor_co_rec_thms, xtor_rel_co_induct, dtor_set_inducts, xtor_co_rec_transfers, xtor_co_rec_o_maps, ...}, lthy)) = fixpoint_bnf false I (construct_fp mixfixes map_bs rel_bs pred_bs set_bss) fp_bs (map dest_TFree As) (map dest_TFree killed_As) (map dest_TFree Xs) ctrXs_repTs empty_comp_cache lthy handle BAD_DEAD (X, X_backdrop) => (case X_backdrop of Type (bad_tc, _) => let val fake_T = qsoty (unfreeze_fp X); val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of 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 (Old_Datatype_Data.get_info (Proof_Context.theory_of 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 " ^ fake_T_backdrop ^ register_hint ()) else error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) end); val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_nesting_rel_eq_onps = map rel_eq_onp_of_bnf fp_nesting_bnfs; val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val live_nesting_rel_eq_onps = map rel_eq_onp_of_bnf live_nesting_bnfs; val liveness = liveness_of_fp_bnf num_As any_fp_bnf; val live = live_of_bnf any_fp_bnf; val _ = if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs @ pred_bs) then warning "Map function, relator, and predicator names ignored" else (); val Bs = @{map 3} (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree_or_tvar S B else A) liveness As Bs0; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val live_AsBs = filter (op <>) (As ~~ Bs); val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_injects = map #abs_inject absT_infos; val abs_inverses = map #abs_inverse absT_infos; val type_definitions = map #type_definition absT_infos; val ctors = map (mk_ctor As) ctors0; val dtors = map (mk_dtor As) dtors0; val fpTs = map (domain_type o fastype_of) dtors; val fpBTs = map B_ify_T fpTs; val real_unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fpTs); val ctr_Tsss = map (map (map real_unfreeze_fp)) ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun define_ctrs_dtrs_for_type_etc fp_bnf fp_b fpT C E ctor dtor xtor_co_rec ctor_dtor dtor_ctor ctor_inject pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm n ks ms abs abs_inject type_definition ctr_bindings ctr_mixfixes ctr_Tss disc_bindings sel_bindingss raw_sel_default_eqs lthy = let val fp_b_name = Binding.name_of fp_b; val ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) = define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy; val ctrs = map (mk_ctr As) ctrs0; val sel_default_eqs = let val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss; val sel_bTs = flat sel_bindingss ~~ flat sel_Tss |> filter_out (Binding.is_empty o fst) |> distinct (Binding.eq_name o apply2 fst); val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy in map (prepare_term sel_default_lthy) raw_sel_default_eqs end; fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b); fun massage_res (ctr_sugar, maps_sets_rels) = (maps_sets_rels, (ctrs, xss, ctor_iff_dtor_thm, ctr_defs, ctr_sugar)); in (wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs #> (fn (ctr_sugar, lthy) => derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps [] fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm [] [] [] ctr_Tss ctr_sugar lthy |>> pair ctr_sugar) ##>> (if fp = Least_FP then define_rec (the recs_args_types) mk_binding fpTs Cs reps else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec #>> apfst massage_res, lthy) end; fun wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types (wrap_one_etcs, lthy) = fold_map I wrap_one_etcs lthy |>> apsnd split_list o apfst (apsnd @{split_list 5} o apfst @{split_list 17} o split_list) o split_list; fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects rel_distincts set_thmss = injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @ set_thmss; fun mk_co_rec_transfer_goals lthy co_recs = let val BE_ify = Term.subst_atomic_types (live_AsBs @ (Cs ~~ Es)); val ((Rs, Ss), names_lthy) = lthy |> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> mk_Frees "S" (map2 mk_pred2T Cs Es); val co_recBs = map BE_ify co_recs; in (Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy) end; fun derive_rec_transfer_thms lthy recs rec_defs (SOME (_, _, _, xsssss)) = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy recs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced nn end; fun derive_rec_o_map_thmss lthy recs rec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_gs = AList.find (op =) (fs ~~ liveness) true; val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs)); val num_rec_args = length rec_arg_Ts; val g_Ts = map B_ify_T rec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names g_Ts; val grecs = map (fn recx => Term.list_comb (Term.map_types B_ify_T recx, gs)) recs; val rec_o_map_lhss = map2 (curry HOLogic.mk_comp) grecs gmaps; val ABfs = (As ~~ Bs) ~~ fs; fun mk_rec_arg_arg (x as Free (_, T)) = let val U = B_ify_T T in if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x end; fun mk_rec_o_map_arg rec_arg_T h = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val xs' = map mk_rec_arg_arg xs; in fold_rev Term.lambda xs (Term.list_comb (h, xs')) end; fun mk_rec_o_map_rhs recx = let val args = map2 mk_rec_o_map_arg rec_arg_Ts gs in Term.list_comb (recx, args) end; val rec_o_map_rhss = map mk_rec_o_map_rhs recs; val rec_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss; val rec_o_map_thms = @{map 3} (fn goal => fn rec_def => fn ctor_rec_o_map => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses ctor_rec_o_map) |> Thm.close_derivation \<^here>) rec_o_map_goals rec_defs xtor_co_rec_o_maps; in map single rec_o_map_thms end; fun derive_note_induct_recs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (ctrss, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (recs, rec_defs)), lthy) = let val ((induct_thms, induct_thm, induct_attrs), (rec_thmss, rec_attrs)) = derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy; val rec_transfer_thmss = map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types); - val induct_type_attr = Attrib.internal o K o Induct.induct_type; - val induct_pred_attr = Attrib.internal o K o Induct.induct_pred; + val induct_type_attr = Attrib.internal \<^here> o K o Induct.induct_type; + val induct_pred_attr = Attrib.internal \<^here> o K o Induct.induct_pred; val ((rel_induct_thmss, common_rel_induct_thms), (rel_induct_attrs, common_rel_induct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) = derive_rel_induct_thms_for_types lthy nn fpTs As Bs ctrss ctr_Tsss (map #exhaust ctr_sugars) xtor_rel_co_induct ctr_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs; in ((map single rel_induct_thms, single common_rel_induct_thm), (rel_induct_attrs, rel_induct_attrs)) end; val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (if nn > 1 then [(inductN, [induct_thm], K induct_attrs), (rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]), (recN, rec_thmss, K rec_attrs), (rec_o_mapN, rec_o_map_thmss, K []), (rec_transferN, rec_transfer_thmss, K []), (rel_inductN, rel_induct_thmss, K (rel_induct_attrs @ [induct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational recs (flat rec_thmss) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (flat rec_thmss)) |> Local_Theory.notes (common_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms (fst (dest_Type (hd fpTs)) ^ implode (map (prefix "_") (tl fp_b_names))) inductN |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars recs rec_defs map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn []) (replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss (replicate nn []) (replicate nn []) rec_transfer_thmss common_rel_induct_thms rel_induct_thmss [] (replicate nn []) rec_o_map_thmss end; fun derive_corec_transfer_thms lthy corecs corec_defs = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy corecs; val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs (flat pgss) pss qssss gssss) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; fun derive_map_o_corec_thmss lthy0 lthy2 corecs corec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_fs = AList.find (op =) (fs ~~ liveness) true; val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0; val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs)); val num_rec_args = length corec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names corec_arg_Ts; val gcorecs = map (fn corecx => Term.list_comb (corecx, gs)) corecs; val map_o_corec_lhss = map2 (curry HOLogic.mk_comp) fmaps gcorecs; val ABfs = (As ~~ Bs) ~~ fs; fun mk_map_o_corec_arg corec_argB_T g = let val T = range_type (fastype_of g); val U = range_type corec_argB_T; in if T = U then g else HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g) end; fun mk_map_o_corec_rhs corecx = let val args = map2 (mk_map_o_corec_arg o B_ify_T) corec_arg_Ts gs in Term.list_comb (B_ify corecx, args) end; val map_o_corec_rhss = map mk_map_o_corec_rhs corecs; val map_o_corec_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) map_o_corec_lhss map_o_corec_rhss; val map_o_corec_thms = @{map 3} (fn goal => fn corec_def => fn dtor_map_o_corec => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt corec_def pre_map_defs live_nesting_map_ident0s abs_inverses dtor_map_o_corec) |> Thm.close_derivation \<^here>) map_o_corec_goals corec_defs xtor_co_rec_o_maps; in map single map_o_corec_thms end; fun derive_note_coinduct_corecs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (_, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (corecs, corec_defs)), lthy) = let val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)], (coinduct_attrs, common_coinduct_attrs)), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) = derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs; fun distinct_prems ctxt thm = Goal.prove (*no sorry*) ctxt [] [] (thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies) (fn _ => HEADGOAL (cut_tac thm THEN' assume_tac ctxt) THEN ALLGOALS eq_assume_tac); fun eq_ifIN _ [thm] = thm | eq_ifIN ctxt (thm :: thms) = distinct_prems ctxt (@{thm eq_ifI} OF (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]}) [thm, eq_ifIN ctxt thms])); val corec_code_thms = map (eq_ifIN lthy) corec_thmss; val corec_sel_thmss = map flat corec_sel_thmsss; - val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type; - val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred; + val coinduct_type_attr = Attrib.internal \<^here> o K o Induct.coinduct_type; + val coinduct_pred_attr = Attrib.internal \<^here> o K o Induct.coinduct_pred; val flat_corec_thms = append oo append; val corec_transfer_thmss = map single (derive_corec_transfer_thms lthy corecs corec_defs); val ((rel_coinduct_thmss, common_rel_coinduct_thms), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_coinduct_thms, common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = derive_rel_coinduct_thms_for_types lthy nn fpTs ns As Bs mss ctr_sugars abs_inverses abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss xtor_rel_co_induct live_nesting_rel_eqs; in ((map single rel_coinduct_thms, single common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) end; val map_o_corec_thmss = derive_map_o_corec_thmss lthy lthy corecs corec_defs; val (set_induct_thms, set_induct_attrss) = derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars) (map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_inducts (map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) dtor_ctors abs_inverses |> split_list; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars (@{map 3} flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss) map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (set_inductN, set_induct_thms, nth set_induct_attrss) :: (if nn > 1 then [(coinductN, [coinduct_thm], K common_coinduct_attrs), (coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs), (rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(coinductN, map single coinduct_thms, fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]), (coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs), (corecN, corec_thmss, K []), (corec_codeN, map single corec_code_thms, K (nitpicksimp_attrs)), (corec_discN, corec_disc_thmss, K []), (corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs), (corec_selN, corec_sel_thmss, K corec_sel_attrs), (corec_transferN, corec_transfer_thmss, K []), (map_o_corecN, map_o_corec_thmss, K []), (rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> fold (Spec_Rules.add Binding.empty Spec_Rules.equational corecs) [flat corec_sel_thmss, flat corec_thmss] |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) corec_code_thms) |> Local_Theory.notes (common_notes @ notes) |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars corecs corec_defs map_thmss [coinduct_thm, coinduct_strong_thm] (transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss corec_disc_iff_thmss (map single corec_code_thms) corec_transfer_thmss common_rel_coinduct_thms rel_coinduct_thmss set_induct_thms (replicate nn (if nn = 1 then set_induct_thms else [])) map_o_corec_thmss end; val lthy = lthy |> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs |> @{fold_map 29} define_ctrs_dtrs_for_type_etc fp_bnfs fp_bs fpTs Cs Es ctors dtors xtor_co_recs ctor_dtors dtor_ctors ctor_injects pre_map_defs pre_set_defss pre_rel_defs xtor_maps xtor_setss xtor_rels ns kss mss abss abs_injects type_definitions ctr_bindingss ctr_mixfixess ctr_Tsss disc_bindingss sel_bindingsss raw_sel_default_eqss |> wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types |> case_fp fp derive_note_induct_recs_thms_for_types derive_note_coinduct_corecs_thms_for_types; val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^ co_prefix fp ^ "datatype")); in lthy end; fun co_datatypes fp = define_co_datatypes (K I) (K I) (K I) (K I) fp; fun co_datatype_cmd fp construct_fp options lthy = define_co_datatypes Plugin_Name.make_filter Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term fp construct_fp options lthy handle EMPTY_DATATYPE s => error ("Cannot define empty datatype " ^ quote s); val parse_ctr_arg = \<^keyword>\(\ |-- parse_binding_colon -- Parse.typ --| \<^keyword>\)\ || Parse.typ >> pair Binding.empty; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.binding parse_ctr_arg -- Parse.opt_mixfix); val parse_spec = parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- parse_ctr_specs) -- parse_map_rel_pred_bindings -- parse_sel_default_eqs; val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec; fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp; end; diff --git a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML @@ -1,519 +1,519 @@ (* Title: HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2013 Suggared flattening of nested to mutual (co)recursion. *) signature BNF_FP_N2M_SUGAR = sig val unfold_lets_splits: term -> term val unfold_splits_lets: term -> term val dest_map: Proof.context -> string -> term -> term * term list val dest_pred: Proof.context -> string -> term -> term * term list val mutualize_fp_sugars: (string -> bool) -> BNF_Util.fp_kind -> int list -> binding list -> typ list -> term list -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory -> (BNF_FP_Def_Sugar.fp_sugar list * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option)) * local_theory val nested_to_mutual_fps: (string -> bool) -> BNF_Util.fp_kind -> binding list -> typ list -> term list -> (term * term list list) list list -> local_theory -> (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option)) * local_theory end; structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR = struct open Ctr_Sugar open BNF_Util open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_FP_N2M val n2mN = "n2m_"; type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option); structure Data = Generic_Data ( type T = n2m_sugar Typtab.table; val empty = Typtab.empty; fun merge data : T = Typtab.merge (K true) data; ); fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) = (map (morph_fp_sugar phi) fp_sugars, (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt, Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt)); val transfer_n2m_sugar = morph_n2m_sugar o Morphism.transfer_morphism; fun n2m_sugar_of ctxt = Typtab.lookup (Data.get (Context.Proof ctxt)) #> Option.map (transfer_n2m_sugar (Proof_Context.theory_of ctxt)); fun register_n2m_sugar key n2m_sugar = - Local_Theory.declaration {syntax = false, pervasive = false} + Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Data.map (Typtab.update (key, morph_n2m_sugar phi n2m_sugar))); fun unfold_lets_splits (Const (\<^const_name>\Let\, _) $ t $ u) = unfold_lets_splits (betapply (unfold_splits_lets u, t)) | unfold_lets_splits (t $ u) = betapply (unfold_lets_splits t, unfold_lets_splits u) | unfold_lets_splits (Abs (s, T, t)) = Abs (s, T, unfold_lets_splits t) | unfold_lets_splits t = t and unfold_splits_lets ((t as Const (\<^const_name>\case_prod\, _)) $ u) = (case unfold_splits_lets u of u' as Abs (s1, T1, Abs (s2, T2, _)) => let val v = Var ((s1 ^ s2, Term.maxidx_of_term u' + 1), HOLogic.mk_prodT (T1, T2)) in lambda v (incr_boundvars 1 (betapplys (u', [HOLogic.mk_fst v, HOLogic.mk_snd v]))) end | _ => t $ unfold_lets_splits u) | unfold_splits_lets (t as Const (\<^const_name>\Let\, _) $ _ $ _) = unfold_lets_splits t | unfold_splits_lets (t $ u) = betapply (unfold_splits_lets t, unfold_lets_splits u) | unfold_splits_lets (Abs (s, T, t)) = Abs (s, T, unfold_splits_lets t) | unfold_splits_lets t = unfold_lets_splits t; fun dest_either_map_or_pred map_or_pred_of_bnf ctxt T_name call = let val bnf = the (bnf_of ctxt T_name); val const0 = map_or_pred_of_bnf bnf; val live = live_of_bnf bnf; val (f_Ts, _) = strip_typeN live (fastype_of const0); val fs = map_index (fn (i, T) => Var (("f", i), T)) f_Ts; val pat = betapplys (const0, fs); val (_, tenv) = fo_match ctxt call pat; in (const0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv []) end; val dest_map = dest_either_map_or_pred map_of_bnf; val dest_pred = dest_either_map_or_pred pred_of_bnf; fun dest_map_or_pred ctxt T_name call = (case try (dest_map ctxt T_name) call of SOME res => res | NONE => dest_pred ctxt T_name call); fun dest_abs_or_applied_map_or_pred _ _ (Abs (_, _, t)) = (Term.dummy, [t]) | dest_abs_or_applied_map_or_pred ctxt s (t1 $ _) = dest_map_or_pred ctxt s t1; fun map_partition f xs = fold_rev (fn x => fn (ys, (good, bad)) => case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad))) xs ([], ([], [])); fun key_of_fp_eqs fp As fpTs Xs ctrXs_repTs = Type (case_fp fp "l" "g", fpTs @ Xs @ ctrXs_repTs) |> Term.map_atyps (fn T as TFree (_, S) => (case find_index (curry (op =) T) As of ~1 => T | j => TFree ("'" ^ string_of_int j, S))); fun get_indices callers t = callers |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE) |> map_filter I; fun mutualize_fp_sugars plugins fp mutual_cliques bs fpTs callers callssss fp_sugars0 no_defs_lthy = let val thy = Proof_Context.theory_of no_defs_lthy; val qsotm = quote o Syntax.string_of_term no_defs_lthy; fun incompatible_calls ts = error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ commas (map qsotm ts)); fun mutual_self_call caller t = error ("Unsupported mutual self-call " ^ qsotm t ^ " from " ^ qsotm caller); fun nested_self_call t = error ("Unsupported nested self-call " ^ qsotm t); val b_names = map Binding.name_of bs; val fp_b_names = map base_name_of_typ fpTs; val nn = length fpTs; val kks = 0 upto nn - 1; fun target_ctr_sugar_of_fp_sugar fpT ({T, fp_ctr_sugar = {ctr_sugar, ...}, ...} : fp_sugar) = let val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) []; val phi = Morphism.term_morphism "BNF" (Term.subst_TVars rho); in morph_ctr_sugar phi ctr_sugar end; val ctor_iff_dtors = map (#ctor_iff_dtor o #fp_ctr_sugar) fp_sugars0; val ctr_defss = map (#ctr_defs o #fp_ctr_sugar) fp_sugars0; val mapss = map (#map_thms o #fp_bnf_sugar) fp_sugars0; val ctr_sugars = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0; val ctrss = map #ctrs ctr_sugars; val ctr_Tss = map (map fastype_of) ctrss; val As' = fold (fold Term.add_tfreesT) ctr_Tss []; val As = map TFree As'; val ((Cs, Xs), _) = no_defs_lthy |> fold Variable.declare_typ As |> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun check_call_dead live_call call = if null (get_indices callers call) then () else incompatible_calls [live_call, call]; fun freeze_fpTs_type_based_default (T as Type (s, Ts)) = (case filter (curry (op =) T o snd) (map_index I fpTs) of [(kk, _)] => nth Xs kk | _ => Type (s, map freeze_fpTs_type_based_default Ts)) | freeze_fpTs_type_based_default T = T; fun freeze_fpTs_mutual_call kk fpT calls T = (case fold (union (op =)) (map (get_indices callers) calls) [] of [] => if T = fpT then nth Xs kk else freeze_fpTs_type_based_default T | [kk'] => if T = fpT andalso kk' <> kk then mutual_self_call (nth callers kk) (the (find_first (not o null o get_indices callers) calls)) else if T = nth fpTs kk' then nth Xs kk' else freeze_fpTs_type_based_default T | _ => incompatible_calls calls); fun freeze_fpTs_map kk (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls)) (Type (s, Ts)) = if Ts' = Ts then nested_self_call live_call else (List.app (check_call_dead live_call) dead_calls; Type (s, map2 (freeze_fpTs_call kk fpT) (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) [] (transpose callss)) Ts)) and freeze_fpTs_call kk fpT calls (T as Type (s, _)) = (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of ([], _) => (case map_partition (try (snd o dest_abs_or_applied_map_or_pred no_defs_lthy s)) calls of ([], _) => freeze_fpTs_mutual_call kk fpT calls T | callsp => freeze_fpTs_map kk fpT callsp T) | callsp => freeze_fpTs_map kk fpT callsp T) | freeze_fpTs_call _ _ _ T = T; val ctr_Tsss = map (map binder_types) ctr_Tss; val ctrXs_Tsss = @{map 4} (map2 o map2 oo freeze_fpTs_call) kks fpTs callssss ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val key = key_of_fp_eqs fp As fpTs Xs ctrXs_repTs; in (case n2m_sugar_of no_defs_lthy key of SOME n2m_sugar => (n2m_sugar, no_defs_lthy) | NONE => let val base_fp_names = Name.variant_list [] fp_b_names; val fp_bs = map2 (fn b_name => fn base_fp_name => Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name))) b_names base_fp_names; val Type (s, Us) :: _ = fpTs; val killed_As' = (case bnf_of no_defs_lthy s of NONE => As' | SOME bnf => let val Type (_, Ts) = T_of_bnf bnf; val deads = deads_of_bnf bnf; val dead_Us = map_filter (fn (T, U) => if member (op =) deads T then SOME U else NONE) (Ts ~~ Us); in fold Term.add_tfreesT dead_Us [] end); val fp_absT_infos = map #absT_info fp_sugars0; val indexed_fp_ress = map (apsnd #fp_res o `(#fp_res_index)) fp_sugars0; val (((pre_bnfs, absT_infos), _), (fp_res as {xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_injects, dtor_ctors, xtor_co_rec_thms, ...}, lthy)) = fixpoint_bnf false I (construct_mutualized_fp fp mutual_cliques fpTs indexed_fp_ress) fp_bs As' killed_As' (map dest_TFree Xs) ctrXs_repTs empty_comp_cache no_defs_lthy; val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_abs_inverses = map #abs_inverse fp_absT_infos; val fp_type_definitions = map #type_definition fp_absT_infos; val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_inverses = map #abs_inverse absT_infos; val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun mk_binding b pre = Binding.prefix_name (pre ^ "_") b; val ((co_recs, co_rec_defs), lthy) = @{fold_map 2} (fn b => if fp = Least_FP then define_rec (the recs_args_types) (mk_binding b) fpTs Cs reps else define_corec (the corecs_args_types) (mk_binding b) fpTs Cs abss) fp_bs xtor_co_recs lthy |>> split_list; val timer = time (timer ("High-level " ^ co_prefix fp ^ "recursors")); val ((common_co_inducts, co_inductss', co_rec_thmss, co_rec_disc_thmss, co_rec_sel_thmsss), fp_sugar_thms) = if fp = Least_FP then derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses fp_type_definitions abs_inverses ctrss ctr_defss co_recs co_rec_defs lthy |> `(fn ((inducts, induct, _), (rec_thmss, _)) => ([induct], [inducts], rec_thmss, replicate nn [], replicate nn [])) ||> (fn info => (SOME info, NONE)) else derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss ctr_sugars co_recs co_rec_defs |> `(fn ((coinduct_thms_pairs, _), corec_thmss, corec_disc_thmss, _, (corec_sel_thmsss, _)) => (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss, corec_disc_thmss, corec_sel_thmsss)) ||> (fn info => (NONE, SOME info)); val timer = time (timer ("High-level " ^ co_prefix fp ^ "induction principles")); val names_lthy = lthy |> fold Variable.declare_typ (As @ Cs @ Xs); val phi = Proof_Context.export_morphism names_lthy lthy; fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctor_iff_dtor ctr_defs ctr_sugar co_rec co_rec_def map_thms co_inducts co_rec_thms co_rec_disc_thms co_rec_sel_thmss ({fp_ctr_sugar = {ctr_transfers, case_transfers, disc_transfers, sel_transfers, ...}, fp_bnf_sugar = {map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases, ...}, fp_co_induct_sugar = SOME {co_rec_disc_iffs, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts, ...}, ...} : fp_sugar) = {T = T, BT = T (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = pre_bnf, fp_bnf = nth (#bnfs fp_res) kk, absT_info = absT_info, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = ctrXs_Tss, ctor_iff_dtor = ctor_iff_dtor, ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, ctr_transfers = ctr_transfers, case_transfers = case_transfers, disc_transfers = disc_transfers, sel_transfers = sel_transfers}, fp_bnf_sugar = {map_thms = map_thms, map_disc_iffs = map_disc_iffs, map_selss = map_selss, rel_injects = rel_injects, rel_distincts = rel_distincts, rel_sels = rel_sels, rel_intros = rel_intros, rel_cases = rel_cases, pred_injects = pred_injects, set_thms = set_thms, set_selssss = set_selssss, set_introssss = set_introssss, set_cases = set_cases}, fp_co_induct_sugar = SOME {co_rec = co_rec, common_co_inducts = common_co_inducts, co_inducts = co_inducts, co_rec_def = co_rec_def, co_rec_thms = co_rec_thms, co_rec_discs = co_rec_disc_thms, co_rec_disc_iffs = co_rec_disc_iffs, co_rec_selss = co_rec_sel_thmss, co_rec_codes = co_rec_codes, co_rec_transfers = co_rec_transfers, co_rec_o_maps = co_rec_o_maps, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = rel_co_inducts, common_set_inducts = common_set_inducts, set_inducts = set_inducts}} |> morph_fp_sugar phi; val target_fp_sugars = @{map 17} mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs mapss (transpose co_inductss') co_rec_thmss co_rec_disc_thmss co_rec_sel_thmsss fp_sugars0; val n2m_sugar = (target_fp_sugars, fp_sugar_thms); in (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar) end) end; fun indexify_callsss ctrs callsss = let fun indexify_ctr ctr = (case AList.lookup Term.aconv_untyped callsss ctr of NONE => replicate (num_binder_types (fastype_of ctr)) [] | SOME callss => map (map (Envir.beta_eta_contract o unfold_lets_splits)) callss); in map indexify_ctr ctrs end; fun retypargs tyargs (Type (s, _)) = Type (s, tyargs); fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) = f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I) | fold_subtype_pairs f TU = f TU; val impossible_caller = Bound ~1; fun nested_to_mutual_fps plugins fp actual_bs actual_Ts actual_callers actual_callssss0 lthy = let val qsoty = quote o Syntax.string_of_typ lthy; val qsotys = space_implode " or " o map qsoty; 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 (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then error (qsoty T ^ " is an old-style datatype") else not_co_datatype0 T | not_co_datatype T = not_co_datatype0 T; fun not_mutually_nested_rec Ts1 Ts2 = error (qsotys Ts1 ^ " is neither mutually " ^ co_prefix fp ^ "recursive with " ^ qsotys Ts2 ^ " nor nested " ^ co_prefix fp ^ "recursive through " ^ (if Ts1 = Ts2 andalso length Ts1 = 1 then "itself" else qsotys Ts2)); val sorted_actual_Ts = sort (prod_ord int_ord Term_Ord.typ_ord o apply2 (`Term.size_of_typ)) actual_Ts; fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s))); fun gen_rhss_in gen_Ts rho (subTs as Type (_, sub_tyargs) :: _) = let fun maybe_insert (T, Type (_, gen_tyargs)) = member (op =) subTs T ? insert (op =) gen_tyargs | maybe_insert _ = I; val gen_ctrs = maps the_ctrs_of gen_Ts; val gen_ctr_Ts = maps (binder_types o fastype_of) gen_ctrs; val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts; in (case fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) [] of [] => [map (typ_subst_nonatomic (map swap rho)) sub_tyargs] | gen_tyargss => gen_tyargss) end; fun the_fp_sugar_of (T as Type (T_name, _)) = (case fp_sugar_of lthy T_name of SOME (fp_sugar as {fp = fp', fp_co_induct_sugar = SOME _, ...}) => if fp = fp' then fp_sugar else not_co_datatype T | _ => not_co_datatype T); fun gather_types _ _ rev_seens gen_seen [] = (rev rev_seens, gen_seen) | gather_types lthy rho rev_seens gen_seen ((T as Type (_, tyargs)) :: Ts) = let val {T = Type (_, tyargs0), fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T; val mutual_Ts = map (retypargs tyargs) mutual_Ts0; val rev_seen = flat rev_seens; val _ = null rev_seens orelse exists (exists_strict_subtype_in rev_seen) mutual_Ts orelse not_mutually_nested_rec mutual_Ts rev_seen; fun fresh_tyargs () = let val (unsorted_gen_tyargs, lthy') = variant_tfrees (replicate (length tyargs) "z") lthy |>> map Logic.varifyT_global; val gen_tyargs = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) tyargs0 unsorted_gen_tyargs; val rho' = (gen_tyargs ~~ tyargs) @ rho; in (rho', gen_tyargs, gen_seen, lthy') end; val (rho', gen_tyargs, gen_seen', lthy') = if exists (exists_subtype_in (flat rev_seens)) mutual_Ts then (case gen_rhss_in gen_seen rho mutual_Ts of [] => fresh_tyargs () | gen_tyargs :: gen_tyargss_tl => let val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl); val mgu = Type.raw_unifys unify_pairs Vartab.empty; val gen_tyargs' = map (Envir.norm_type mgu) gen_tyargs; val gen_seen' = map (Envir.norm_type mgu) gen_seen; in (rho, gen_tyargs', gen_seen', lthy) end) else fresh_tyargs (); val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0; val other_mutual_Ts = remove1 (op =) T mutual_Ts; val Ts' = fold (remove1 (op =)) other_mutual_Ts Ts; in gather_types lthy' rho' (mutual_Ts :: rev_seens) (gen_seen' @ gen_mutual_Ts) Ts' end | gather_types _ _ _ _ (T :: _) = not_co_datatype T; val (perm_Tss, perm_gen_Ts) = gather_types lthy [] [] [] sorted_actual_Ts; val (perm_mutual_cliques, perm_Ts) = split_list (flat (map_index (fn (i, perm_Ts) => map (pair i) perm_Ts) perm_Tss)); val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts; val missing_Ts = fold (remove1 (op =)) actual_Ts perm_Ts; val Ts = actual_Ts @ missing_Ts; val nn = length Ts; val kks = 0 upto nn - 1; val callssss0 = pad_list [] nn actual_callssss0; val common_name = mk_common_name (map Binding.name_of actual_bs); val bs = pad_list (Binding.name common_name) nn actual_bs; val callers = pad_list impossible_caller nn actual_callers; fun permute xs = permute_like (op =) Ts perm_Ts xs; fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs; (* The assignment of callers to mutual cliques is incorrect in general. This code would need to inspect the actual calls to discover the correct cliques in the presence of type duplicates. However, the naive scheme implemented here supports "prim(co)rec" specifications following reasonable ordering of the duplicates (e.g., keeping the cliques together). *) val perm_bs = permute bs; val perm_callers = permute callers; val perm_kks = permute kks; val perm_callssss0 = permute callssss0; val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts; val perm_callssss = map2 (indexify_callsss o #ctrs o #ctr_sugar o #fp_ctr_sugar) perm_fp_sugars0 perm_callssss0; val ((perm_fp_sugars, fp_sugar_thms), lthy) = if length perm_Tss = 1 then ((perm_fp_sugars0, (NONE, NONE)), lthy) else mutualize_fp_sugars plugins fp perm_mutual_cliques perm_bs perm_frozen_gen_Ts perm_callers perm_callssss perm_fp_sugars0 lthy; val fp_sugars = unpermute perm_fp_sugars; in ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy) end; end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec.ML b/src/HOL/Tools/BNF/bnf_gfp_grec.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML @@ -1,3228 +1,3228 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec.ML Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Aymeric Bouzy, Ecole polytechnique Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor construction. *) signature BNF_GFP_GREC = sig val Tsubst: typ -> typ -> typ -> typ val substT: typ -> typ -> term -> term val freeze_types: Proof.context -> (indexname * sort) list -> typ list -> typ list val dummify_atomic_types: term -> term val define_const: bool -> binding -> int -> string -> term -> local_theory -> (term * thm) * local_theory type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table} val map_buffer: (term -> term) -> buffer -> buffer val specialize_buffer_types: buffer -> buffer type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list} type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: BNF_FP_Def_Sugar.fp_sugar list, ssig_fp_sugar: BNF_FP_Def_Sugar.fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: BNF_Def.bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info} type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm} val not_codatatype: Proof.context -> typ -> 'a val mk_fp_binding: binding -> string -> binding val bnf_kill_all_but: int -> BNF_Def.bnf -> local_theory -> BNF_Def.bnf * local_theory val print_corec_infos: Proof.context -> unit val has_no_corec_info: Proof.context -> string -> bool val corec_info_of: typ -> local_theory -> corec_info * local_theory val maybe_corec_info_of: Proof.context -> typ -> corec_info option val corec_infos_of: Proof.context -> string -> corec_info list val corec_infos_of_generic: Context.generic -> Symtab.key -> corec_info list val prepare_friend_corec: string -> typ -> local_theory -> (corec_info * binding * int * typ * typ * typ * typ * typ * BNF_Def.bnf * BNF_Def.bnf * BNF_FP_Def_Sugar.fp_sugar * BNF_FP_Def_Sugar.fp_sugar * buffer) * local_theory val register_friend_corec: string -> binding -> int -> typ -> typ -> typ -> BNF_Def.bnf -> BNF_FP_Def_Sugar.fp_sugar -> BNF_FP_Def_Sugar.fp_sugar -> term -> term -> thm -> corec_info -> local_theory -> friend_info * local_theory end; structure BNF_GFP_Grec : BNF_GFP_GREC = struct open Ctr_Sugar open BNF_Util open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_LFP open BNF_FP_Def_Sugar open BNF_LFP_Rec_Sugar open BNF_GFP_Grec_Tactics val algLamN = "algLam"; val algLam_algLamN = "algLam_algLam"; val algLam_algrhoN = "algLam_algrho"; val algrhoN = "algrho"; val CLeafN = "CLeaf"; val congN = "congclp"; val cong_alg_introsN = "cong_alg_intros"; val cong_localeN = "cong_locale"; val corecUUN = "corecUU"; val corecUU_transferN = "corecUU_transfer"; val corecUU_uniqueN = "corecUU_unique"; val cutSsigN = "cutSsig"; val dtor_algLamN = "dtor_algLam"; val dtor_algrhoN = "dtor_algrho"; val dtor_coinductN = "dtor_coinduct"; val dtor_transferN = "dtor_transfer"; val embLN = "embL"; val embLLN = "embLL"; val embLRN = "embLR"; val embL_pointful_naturalN = "embL_pointful_natural"; val embL_transferN = "embL_transfer"; val equivp_RetrN = "equivp_Retr"; val evalN = "eval"; val eval_coreN = "eval_core"; val eval_core_pointful_naturalN = "eval_core_pointful_natural"; val eval_core_transferN = "eval_core_transfer"; val eval_flatN = "eval_flat"; val eval_simpsN = "eval_simps"; val flatN = "flat"; val flat_pointful_naturalN = "flat_pointful_natural"; val flat_transferN = "flat_transfer"; val k_as_ssig_naturalN = "k_as_ssig_natural"; val k_as_ssig_transferN = "k_as_ssig_transfer"; val LamN = "Lam"; val Lam_transferN = "Lam_transfer"; val Lam_pointful_naturalN = "Lam_pointful_natural"; val OperN = "Oper"; val proto_sctrN = "proto_sctr"; val proto_sctr_pointful_naturalN = "proto_sctr_pointful_natural"; val proto_sctr_transferN = "proto_sctr_transfer"; val rho_transferN = "rho_transfer"; val Retr_coinductN = "Retr_coinduct"; val sctrN = "sctr"; val sctr_transferN = "sctr_transfer"; val sctr_pointful_naturalN = "sctr_pointful_natural"; val sigN = "sig"; val SigN = "Sig"; val Sig_pointful_naturalN = "Sig_pointful_natural"; val corecUN = "corecU"; val corecU_ctorN = "corecU_ctor"; val corecU_uniqueN = "corecU_unique"; val unsigN = "unsig"; val VLeafN = "VLeaf"; val s_prefix = "s"; (* transforms "sig" into "ssig" *) fun not_codatatype ctxt T = error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T); fun mutual_codatatype () = error ("Mutually corecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\primcorec\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun noncorecursive_codatatype () = error ("Noncorecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\definition\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun singleton_codatatype ctxt = error ("Singleton corecursive codatatypes are not supported (use " ^ quote (Syntax.string_of_typ ctxt \<^typ>\unit\) ^ " instead)"); fun merge_lists eq old1 old2 = (old1 |> subtract eq old2) @ old2; fun add_type_namesT (Type (s, Ts)) = insert (op =) s #> fold add_type_namesT Ts | add_type_namesT _ = I; fun Tsubst Y T = Term.typ_subst_atomic [(Y, T)]; fun substT Y T = Term.subst_atomic_types [(Y, T)]; fun freeze_types ctxt except_tvars Ts = let val As = fold Term.add_tvarsT Ts [] |> subtract (op =) except_tvars; val (Bs, _) = ctxt |> mk_TFrees' (map snd As); in map (Term.typ_subst_TVars (map fst As ~~ Bs)) Ts end; fun typ_unify_disjointly thy (T, T') = if T = T' then T else let val tvars = Term.add_tvar_namesT T []; val tvars' = Term.add_tvar_namesT T' []; val maxidx' = maxidx_of_typ T'; val T = T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1); val maxidx = Integer.max (maxidx_of_typ T) maxidx'; val (tyenv, _) = Sign.typ_unify thy (T, T') (Vartab.empty, maxidx); in Envir.subst_type tyenv T end; val dummify_atomic_types = Term.map_types (Term.map_atyps (K Term.dummyT)); fun mk_internal internal ctxt f = if internal andalso not (Config.get ctxt bnf_internals) then f else I fun mk_fp_binding fp_b pre = Binding.map_name (K pre) fp_b |> Binding.qualify true (Binding.name_of fp_b); fun mk_version_binding version = Binding.qualify false ("v" ^ string_of_int version); fun mk_version_fp_binding internal ctxt = mk_internal internal ctxt Binding.concealed ooo (mk_fp_binding oo mk_version_binding); (*FIXME: get rid of ugly names when typedef and primrec respect qualification*) fun mk_version_binding_ugly version = Binding.suffix_name ("_v" ^ string_of_int version); fun mk_version_fp_binding_ugly internal ctxt version fp_b pre = Binding.prefix_name (pre ^ "_") fp_b |> mk_version_binding_ugly version |> mk_internal internal ctxt Binding.concealed; fun mk_mapN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (TA --> TB) (map_of_bnf bnf) end; fun mk_relN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (mk_pred2T TA TB) (rel_of_bnf bnf) end; fun mk_map1 ctxt Y Z = mk_mapN ctxt [(Y, Z)]; fun mk_rel1 ctxt Y Z = mk_relN ctxt [(Y, Z)]; fun define_const internal fp_b version name rhs lthy = let val b = mk_version_fp_binding internal lthy version fp_b name; val ((free, (_, def_free)), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((b, NoSyn), ((Thm.def_binding b |> Binding.concealed, []), rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free), lthy) end; fun define_single_primrec b eqs lthy = let val (([free], [def_free], [simps_free]), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> primrec false [] [(b, NONE, NoSyn)] (map (fn eq => ((Binding.empty_atts, eq), [], [])) eqs) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free, map (Morphism.thm phi) simps_free), lthy) end; type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table}; fun map_buffer f {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = {Oper = f Oper, VLeaf = f VLeaf, CLeaf = f CLeaf, ctr_wrapper = f ctr_wrapper, friends = Symtab.map (K (apsnd f)) friends}; fun morph_buffer phi = map_buffer (Morphism.term phi); fun specialize_buffer_types {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = let val ssig_T as Type (_, Ts) = body_type (fastype_of VLeaf); val Y = List.last Ts; val ssigifyT = substT Y ssig_T; in {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ssigifyT ctr_wrapper, friends = Symtab.map (K (apsnd ssigifyT)) friends} end; type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list}; fun map_dtor_coinduct_info f {dtor_coinduct, cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros} = {dtor_coinduct = f dtor_coinduct, cong_def = f cong_def, cong_locale = f cong_locale, cong_base = f cong_base, cong_refl = f cong_refl, cong_sym = f cong_sym, cong_trans = f cong_trans, cong_alg_intros = map f cong_alg_intros}; fun morph_dtor_coinduct_info phi = map_dtor_coinduct_info (Morphism.thm phi); type corec_ad = {fpT: typ, friend_names: string list}; fun morph_corec_ad phi {fpT, friend_names} = {fpT = Morphism.typ phi fpT, friend_names = friend_names}; type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: fp_sugar list, ssig_fp_sugar: fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info}; fun morph_corec_info phi ({fp_b, version, fpT, Y, Z, friend_names, sig_fp_sugars, ssig_fp_sugar, Lam, proto_sctr, flat, eval_core, eval, algLam, corecUU, dtor_transfer, Lam_transfer, Lam_pointful_natural, proto_sctr_transfer, flat_simps, eval_core_simps, eval_thm, eval_simps, all_algLam_algs, algLam_thm, dtor_algLam, corecUU_thm, corecUU_unique, corecUU_transfer, buffer, all_dead_k_bnfs, Retr, equivp_Retr, Retr_coinduct, dtor_coinduct_info} : corec_info) = {fp_b = fp_b, version = version, fpT = Morphism.typ phi fpT, Y = Morphism.typ phi Y, Z = Morphism.typ phi Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugars (*no morph*), ssig_fp_sugar = ssig_fp_sugar (*no morph*), Lam = Morphism.term phi Lam, proto_sctr = Morphism.term phi proto_sctr, flat = Morphism.term phi flat, eval_core = Morphism.term phi eval_core, eval = Morphism.term phi eval, algLam = Morphism.term phi algLam, corecUU = Morphism.term phi corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Morphism.thm phi Lam_transfer, Lam_pointful_natural = Morphism.thm phi Lam_pointful_natural, proto_sctr_transfer = Morphism.thm phi proto_sctr_transfer, flat_simps = map (Morphism.thm phi) flat_simps, eval_core_simps = map (Morphism.thm phi) eval_core_simps, eval_thm = Morphism.thm phi eval_thm, eval_simps = map (Morphism.thm phi) eval_simps, all_algLam_algs = map (Morphism.thm phi) all_algLam_algs, algLam_thm = Morphism.thm phi algLam_thm, dtor_algLam = Morphism.thm phi dtor_algLam, corecUU_thm = Morphism.thm phi corecUU_thm, corecUU_unique = Morphism.thm phi corecUU_unique, corecUU_transfer = Morphism.thm phi corecUU_transfer, buffer = morph_buffer phi buffer, all_dead_k_bnfs = map (morph_bnf phi) all_dead_k_bnfs, Retr = Morphism.term phi Retr, equivp_Retr = Morphism.thm phi equivp_Retr, Retr_coinduct = Morphism.thm phi Retr_coinduct, dtor_coinduct_info = morph_dtor_coinduct_info phi dtor_coinduct_info}; datatype ('a, 'b) expr = Ad of 'a * (local_theory -> 'b * local_theory) | Info of 'b; fun is_Ad (Ad _) = true | is_Ad _ = false; fun is_Info (Info _) = true | is_Info _ = false; type corec_info_expr = (corec_ad, corec_info) expr; fun morph_corec_info_expr phi (Ad (ad, f)) = Ad (morph_corec_ad phi ad, f) | morph_corec_info_expr phi (Info info) = Info (morph_corec_info phi info); val transfer_corec_info_expr = morph_corec_info_expr o Morphism.transfer_morphism; type corec_data = int Symtab.table * corec_info_expr list Symtab.table list; structure Data = Generic_Data ( type T = corec_data; val empty = (Symtab.empty, [Symtab.empty]); fun merge ((version_tab1, info_tabs1), (version_tab2, info_tabs2)) : T = (Symtab.join (K Int.max) (version_tab1, version_tab2), info_tabs1 @ info_tabs2); ); fun corec_ad_of_expr (Ad (ad, _)) = ad | corec_ad_of_expr (Info {fpT, friend_names, ...}) = {fpT = fpT, friend_names = friend_names}; fun corec_info_exprs_of_generic context fpT_name = let val thy = Context.theory_of context; val info_tabs = snd (Data.get context); in maps (fn info_tab => these (Symtab.lookup info_tab fpT_name)) info_tabs |> map (transfer_corec_info_expr thy) end; val corec_info_exprs_of = corec_info_exprs_of_generic o Context.Proof; val keep_corec_infos = map_filter (fn Ad _ => NONE | Info info => SOME info); val corec_infos_of_generic = keep_corec_infos oo corec_info_exprs_of_generic; val corec_infos_of = keep_corec_infos oo corec_info_exprs_of; fun str_of_corec_ad ctxt {fpT, friend_names} = "[" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "]"; fun str_of_corec_info ctxt {fpT, version, friend_names, ...} = "{" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "; v" ^ string_of_int version ^ "}"; fun str_of_corec_info_expr ctxt (Ad (ad, _)) = str_of_corec_ad ctxt ad | str_of_corec_info_expr ctxt (Info info) = str_of_corec_info ctxt info; fun print_corec_infos ctxt = Symtab.fold (fn (fpT_name, exprs) => fn () => writeln (fpT_name ^ ":\n" ^ cat_lines (map (prefix " " o str_of_corec_info_expr ctxt) exprs))) (the_single (snd (Data.get (Context.Proof ctxt)))) (); val has_no_corec_info = null oo corec_info_exprs_of; fun get_name_next_version_of fpT_name ctxt = let val (version_tab, info_tabs) = Data.get (Context.Theory (Proof_Context.theory_of ctxt)); val fp_base = Long_Name.base_name fpT_name; val fp_b = Binding.name fp_base; val version_tab' = Symtab.map_default (fp_base, ~1) (Integer.add 1) version_tab; val SOME version = Symtab.lookup version_tab' fp_base; val ctxt' = ctxt |> Local_Theory.background_theory (Context.theory_map (Data.put (version_tab', info_tabs))); in ((fp_b, version), ctxt') end; type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm}; fun morph_friend_info phi ({algrho, dtor_algrho, algLam_algrho} : friend_info) = {algrho = Morphism.term phi algrho, dtor_algrho = Morphism.thm phi dtor_algrho, algLam_algrho = Morphism.thm phi algLam_algrho}; fun checked_fp_sugar_of ctxt fpT_name = let val fp_sugar as {X, fp_res = {Ts = fpTs, ...}, fp_ctr_sugar = {ctrXs_Tss, ...}, ...} = (case fp_sugar_of ctxt fpT_name of SOME (fp_sugar as {fp = Greatest_FP, ...}) => fp_sugar | _ => not_codatatype ctxt (Type (fpT_name, [] (*yuck*)))); val _ = if length fpTs > 1 then mutual_codatatype () else if not (exists (exists (Term.exists_subtype (curry (op =) X))) ctrXs_Tss) then noncorecursive_codatatype () else if ctrXs_Tss = [[X]] then singleton_codatatype ctxt else (); in fp_sugar end; fun bnf_kill_all_but nn bnf lthy = ((empty_comp_cache, empty_unfolds), lthy) |> kill_bnf I (live_of_bnf bnf - nn) bnf ||> snd; fun bnf_with_deads_and_lives dead_Es live_As Y fpT T lthy = let val qsoty = quote o Syntax.string_of_typ lthy; val unfreeze_fp = Tsubst Y fpT; fun flatten_tyargs Ass = map dest_TFree live_As |> filter (fn T => exists (fn Ts => member (op =) Ts T) Ass); val ((bnf, _), (_, lthy)) = bnf_of_typ false Do_Inline I flatten_tyargs [Term.dest_TFree Y] (map Term.dest_TFree dead_Es) T ((empty_comp_cache, empty_unfolds), lthy) handle BAD_DEAD (Y, Y_backdrop) => (case Y_backdrop of Type (bad_tc, _) => let val T = qsoty (unfreeze_fp Y); val T_backdrop = qsoty (unfreeze_fp Y_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible occurrence of type " ^ T ^ " in type expression " ^ T_backdrop) else error ("Unsupported occurrence of type " ^ T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ T_backdrop ^ register_hint ()) end); val phi = Morphism.term_morphism "BNF" (Raw_Simplifier.rewrite_term (Proof_Context.theory_of lthy) @{thms BNF_Composition.id_bnf_def} []) $> Morphism.thm_morphism "BNF" (unfold_thms lthy @{thms BNF_Composition.id_bnf_def}); in (morph_bnf phi bnf, lthy) end; fun define_sig_type fp_b version fp_alives Es Y rhsT lthy = let val T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val ctr_b = mk_version_fp_binding false lthy version fp_b SigN; val sel_b = mk_version_fp_binding true lthy version fp_b unsigN; val lthy = (snd o Local_Theory.begin_nested) lthy; val T_name = Local_Theory.full_name lthy T_b; val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, ctr_b), [(sel_b, rhsT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = true; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.end_nested; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun define_ssig_type fp_b version fp_alives Es Y fpT lthy = let val sig_T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val T_b = Binding.prefix_name s_prefix sig_T_b; val Oper_b = mk_version_fp_binding false lthy version fp_b OperN; val VLeaf_b = mk_version_fp_binding false lthy version fp_b VLeafN; val CLeaf_b = mk_version_fp_binding false lthy version fp_b CLeafN; val lthy = (snd o Local_Theory.begin_nested) lthy; val sig_T_name = Local_Theory.full_name lthy sig_T_b; val T_name = Long_Name.map_base_name (prefix s_prefix) sig_T_name; val As = Es @ [Y]; val ssig_sig_T = Type (sig_T_name, Es @ [Type (T_name, As)]); val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, Oper_b), [(Binding.empty, ssig_sig_T)]), NoSyn), (((Binding.empty, VLeaf_b), [(Binding.empty, Y)]), NoSyn), (((Binding.empty, CLeaf_b), [(Binding.empty, fpT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = false; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.end_nested; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun embed_Sig ctxt Sig inl_or_r t = Library.foldl1 HOLogic.mk_comp [Sig, inl_or_r, dummify_atomic_types t] |> Syntax.check_term ctxt; fun mk_ctr_wrapper_friends ctxt friend_name friend_T old_sig_T k_T Sig old_buffer = let val embed_Sig_inl = embed_Sig ctxt Sig (Inl_const old_sig_T k_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old_buffer); val friends = Symtab.map (K (apsnd embed_Sig_inl)) (#friends old_buffer) |> Symtab.update_new (friend_name, (friend_T, HOLogic.mk_comp (Sig, Inr_const old_sig_T k_T))); in (ctr_wrapper, friends) end; fun pre_type_of_ctor Y ctor = let val (fp_preT, fpT) = dest_funT (fastype_of ctor); in typ_subst_nonatomic [(fpT, Y)] fp_preT end; fun mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf = let val inr' = Inr_const old_sig_T k_T; val dead_sig_map' = substT Z ssig_T dead_sig_map; in Library.foldl1 HOLogic.mk_comp [Oper, dead_sig_map' $ VLeaf, Sig, inr'] end; fun define_embL name fp_b version Y Z fpT old_sig_T old_ssig_T other_summand_T ssig_T Inl_or_r_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf lthy = let val embL_b = mk_version_fp_binding true lthy version fp_b name; val old_ssig_old_sig_T = Tsubst Y old_ssig_T old_sig_T; val ssig_old_sig_T = Tsubst Y ssig_T old_sig_T; val ssig_other_summand_T = Tsubst Y ssig_T other_summand_T; val sigx = Var (("s", 0), old_ssig_old_sig_T); val x = Var (("x", 0), Y); val j = Var (("j", 0), fpT); val embL = Free (Binding.name_of embL_b, old_ssig_T --> ssig_T); val dead_old_sig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_sig_map; val Sig' = substT Y ssig_T Sig; val inl' = Inl_or_r_const ssig_old_sig_T ssig_other_summand_T; val Oper_eq = mk_Trueprop_eq (embL $ (old_Oper $ sigx), Oper $ (Sig' $ (inl' $ (dead_old_sig_map' $ embL $ sigx)))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (embL $ (old_VLeaf $ x), VLeaf $ x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (embL $ (old_CLeaf $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec embL_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val snd' = snd_const YpreT; val dead_pre_map' = substT Z ssig_T dead_pre_map; val Sig' = substT Y ssig_T Sig; val unsig' = substT Y ssig_T unsig; val dead_sig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, ssig_T)] dead_sig_map; val rhs = HOLogic.mk_comp (unsig', dead_sig_map' $ Library.foldl1 HOLogic.mk_comp [Oper, Sig', dead_pre_map' $ VLeaf, snd']); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val unsig' = substT Y YpreT unsig; val rhs = HOLogic.mk_comp (mk_case_sum (left_case, right_case), unsig'); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case rho lthy end; fun define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old1_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, old2_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embLL, old1_Lam); val right_case = HOLogic.mk_comp (dead_pre_map'' $ embLR, old2_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy end; fun define_proto_sctr_step_or_merge fp_b version old_sig_T right_T Sig old_proto_sctr = let val rhs = Library.foldl1 HOLogic.mk_comp [Sig, Inl_const old_sig_T right_T, old_proto_sctr]; in define_const true fp_b version proto_sctrN rhs end; fun define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map lthy = let val flat_b = mk_version_fp_binding true lthy version fp_b flatN; val ssig_sig_T = Tsubst Y ssig_T sig_T; val ssig_ssig_sig_T = Tsubst Y ssig_T ssig_sig_T; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), ssig_ssig_sig_T); val x = Var (("x", 0), ssig_T); val j = Var (("j", 0), fpT); val flat = Free (Binding.name_of flat_b, ssig_ssig_T --> ssig_T); val Oper' = substT Y ssig_T Oper; val VLeaf' = substT Y ssig_T VLeaf; val CLeaf' = substT Y ssig_T CLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_sig_map; val Oper_eq = mk_Trueprop_eq (flat $ (Oper' $ sigx), Oper $ (dead_sig_map' $ flat $ sigx)) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (flat $ (VLeaf' $ x), x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (flat $ (CLeaf' $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec flat_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam lthy = let val eval_core_b = mk_version_fp_binding true lthy version fp_b eval_coreN; val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_sig_T = Tsubst Y Ypre_ssig_T sig_T; val ssig_preT = Tsubst Y ssig_T preT; val ssig_YpreT = Tsubst Y ssig_T YpreT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), Ypre_ssig_sig_T); val x = Var (("x", 0), YpreT); val j = Var (("j", 0), fpT); val eval_core = Free (Binding.name_of eval_core_b, Ypre_ssig_T --> ssig_preT); val Oper' = substT Y YpreT Oper; val VLeaf' = substT Y YpreT VLeaf; val CLeaf' = substT Y YpreT CLeaf; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = substT Z ssig_T dead_pre_map; val dead_pre_map''' = Term.subst_atomic_types [(Y, fpT), (Z, ssig_T)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_sig_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val Lam' = substT Y ssig_T Lam; val fst' = fst_const YpreT; val snd' = snd_const YpreT; val Oper_eq = mk_Trueprop_eq (eval_core $ (Oper' $ sigx), dead_pre_map' $ flat $ (Lam' $ (dead_sig_map' $ (Abs (Name.uu, Ypre_ssig_T, HOLogic.mk_prod (dead_ssig_map' $ fst' $ Bound 0, eval_core $ Bound 0))) $ sigx))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (eval_core $ (VLeaf' $ x), dead_pre_map'' $ VLeaf $ (snd' $ x)) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (eval_core $ (CLeaf' $ j), dead_pre_map''' $ CLeaf $ (dtor $ j)) |> Logic.all j; in define_single_primrec eval_core_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core lthy = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_ssig_T = Tsubst Y fpT ssig_T; val dtor_unfold' = substT Z fp_ssig_T dtor_unfold; val dead_ssig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_ssig_map; val eval_core' = substT Y fpT eval_core; val id' = HOLogic.id_const fpT; val rhs = dtor_unfold' $ HOLogic.mk_comp (eval_core', dead_ssig_map' $ mk_convol (id', dtor)); in define_const true fp_b version evalN rhs lthy end; fun define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val h = Var (("h", 0), Y --> ssig_preT); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = substT Z ssig_ssig_preT dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val rhs = Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (VLeaf, h)] |> Term.lambda h; in define_const true fp_b version cutSsigN rhs lthy end; fun define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval lthy = let val fp_ssig_T = Tsubst Y fpT ssig_T; val Oper' = substT Y fpT Oper; val VLeaf' = substT Y fpT VLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_ssig_T)] dead_sig_map; val rhs = Library.foldl1 HOLogic.mk_comp [eval, Oper', dead_sig_map' $ VLeaf']; in define_const true fp_b version algLamN rhs lthy end; fun define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig lthy = let val ssig_preT = Tsubst Y ssig_T preT; val h = Var (("h", 0), Y --> ssig_preT); val dtor_unfold' = substT Z ssig_T dtor_unfold; val rhs = HOLogic.mk_comp (dtor_unfold' $ (cutSsig $ h), VLeaf) |> Term.lambda h; in define_const true fp_b version corecUN rhs lthy end; fun define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val h = Var (("h", 0), Y --> ssig_pre_ssig_T); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val eval_core' = substT Y ssig_T eval_core; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_ssig_preT)] dead_ssig_map; val id' = HOLogic.id_const ssig_preT; val rhs = corecU $ Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (sctr, id'), h] |> Term.lambda h; in define_const true fp_b version corecUUN rhs lthy end; fun derive_sig_transfer maybe_swap ctxt live_AsBs pre_rel sig_rel Rs R const pre_rel_def preT_rel_eqs transfer_thm = let val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRsig_rel = list_comb (sig_rel, Rs) $ R; val constB = Term.subst_atomic_types live_AsBs const; val goal = uncurry mk_rel_fun (maybe_swap (RRpre_rel, RRsig_rel)) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_sig_transfer_tac ctxt pre_rel_def preT_rel_eqs transfer_thm)) |> Thm.close_derivation \<^here> end; fun derive_transfer_by_transfer_prover ctxt live_AsBs Rs R const const_defs rel_eqs transfers = let val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_parametricity_goal ctxt (Rs @ [R]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt (const_defs @ map (fn thm => thm RS sym) rel_eqs) rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_dtor_transfer ctxt live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm = let val Type (\<^type_name>\fun\, [fpT, Type (\<^type_name>\fun\, [fpTB, \<^typ>\bool\])]) = snd (strip_typeN (length live_EsFs) (fastype_of fp_rel)); val pre_rel' = Term.subst_atomic_types [(Y, fpT), (Z, fpTB)] pre_rel; val Rpre_rel = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val dtorB = Term.subst_atomic_types live_EsFs dtor; val goal = HOLogic.mk_Trueprop (mk_rel_fun Rfp_rel (Rpre_rel $ Rfp_rel) $ dtor $ dtorB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_transfer_tac ctxt dtor_rel_thm)) |> Thm.close_derivation \<^here> end; fun derive_Lam_or_eval_core_transfer ctxt live_AsBs Y Z preT ssig_T Rs R pre_rel sig_or_ssig_rel ssig_rel const const_def rel_eqs transfers = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreTB = typ_subst_atomic live_AsBs YpreT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val sig_or_ssig_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreTB)] sig_or_ssig_rel; val Rsig_or_ssig_rel' = list_comb (sig_or_ssig_rel', Rs); val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rpre_rel' = list_comb (pre_rel', Rs); val constB = subst_atomic_types live_AsBs const; val goal = mk_rel_fun (Rsig_or_ssig_rel' $ mk_rel_prod R RRpre_rel) (Rpre_rel' $ RRssig_rel) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [const_def] rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_proto_sctr_transfer_step_or_merge ctxt Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs transfers = let val proto_sctrZ = substT Y Z proto_sctr; val goal = mk_rel_fun (dead_pre_rel $ R) (dead_sig_rel $ R) $ proto_sctr $ proto_sctrZ |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [proto_sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_sctr_transfer ctxt live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs transfers = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val sctrB = subst_atomic_types live_AsBs sctr; val goal = HOLogic.mk_Trueprop (mk_rel_fun (Rpre_rel' $ RRssig_rel) RRssig_rel $ sctr $ sctrB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_transfer ctxt live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs transfers = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val ssig_preTB = typ_subst_atomic live_AsBs ssig_preT; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val ssig_rel' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_preTB)] ssig_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rssig_rel' = list_comb (ssig_rel', Rs); val corecUUB = subst_atomic_types live_AsBs corecUU; val goal = mk_rel_fun (mk_rel_fun R (Rssig_rel' $ (Rpre_rel' $ RRssig_rel))) (mk_rel_fun R Rfp_rel) $ corecUU $ corecUUB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [cutSsig_def, corecU_def, corecUU_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun mk_natural_goal ctxt simple_T_mapfs fs t u = let fun build_simple (T, _) = (case AList.lookup (op =) simple_T_mapfs T of SOME mapf => mapf | NONE => the (find_first (fn f => domain_type (fastype_of f) = T) fs)); val simple_Ts = map fst simple_T_mapfs; val t_map = build_map ctxt simple_Ts [] build_simple (apply2 (range_type o fastype_of) (t, u)); val u_map = build_map ctxt simple_Ts [] build_simple (apply2 (domain_type o fastype_of) (t, u)); in mk_Trueprop_eq (HOLogic.mk_comp (u, u_map), HOLogic.mk_comp (t_map, t)) end; fun derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f const map_thms = let val ffpre_map = list_comb (pre_map, fs) $ f; val constB = subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt [(preT, ffpre_map)] (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_by_unfolding_tac ctxt map_thms)) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer ctxt live_AsBs simple_T_mapfs fs f const transfer bnfs subst_bnfs = let val m = length live_AsBs; val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt simple_T_mapfs (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_from_transfer_tac ctxt m (replicate m true) transfer [] (map rel_Grp_of_bnf bnfs) (map rel_Grp_of_bnf subst_bnfs))) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val preT' = Term.typ_subst_atomic [(Y, ssig_T), (Z, ssig_TB)] preT; val ffpre_map = list_comb (pre_map, fs) $ f; val pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_map; val fpre_map' = list_comb (pre_map', fs); val ffssig_map = list_comb (ssig_map, fs) $ f; val preT_mapfs = [(preT, ffpre_map), (preT', fpre_map' $ ffssig_map)]; in derive_natural_from_transfer ctxt live_AsBs preT_mapfs fs f end; fun derive_Lam_Inl_Inr ctxt Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_sig_T = Tsubst Y YpreT old_sig_T; val Ypre_k_T = Tsubst Y YpreT k_T; val inl' = Inl_const Ypre_old_sig_T Ypre_k_T; val inr' = Inr_const Ypre_old_sig_T Ypre_k_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val Sig' = substT Y YpreT Sig; val Lam_o_Sig = HOLogic.mk_comp (Lam, Sig'); val inl_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inl'), HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam)); val inr_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inr'), rho); val goals = [inl_goal, inr_goal]; val goal = Logic.mk_conjunction_balanced goals; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_Lam_Inl_Inr_tac ctxt unsig_thm Lam_def)) |> Conjunction.elim_balanced (length goals) |> map (Thm.close_derivation \<^here>) end; fun derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps = let val x' = substT Y ssig_T x; val dead_ssig_map' = substT Z ssig_T dead_ssig_map; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ VLeaf $ x'), x'); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (fp_map_id :: sig_map_ident :: sig_map_comp :: ssig_map_thms @ flat_simps @ @{thms o_apply id_apply id_def[symmetric]}))) |> Thm.close_derivation \<^here> end; fun derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps = let val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_ssig_T = Tsubst Y ssig_T ssig_ssig_T; val x' = substT Y ssig_ssig_ssig_T x; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_ssig_map; val flat' = substT Y ssig_T flat; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ flat $ x'), flat $ (flat' $ x')); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (o_apply :: fp_map_id :: sig_map_comp :: ssig_map_thms @ flat_simps))) |> Thm.close_derivation \<^here> end; fun derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_ssig_T = Tsubst Y YpreT ssig_ssig_T; val ssig_YpreT = Tsubst Y ssig_T YpreT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_ssig_map; val dead_ssig_map'' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val flat' = substT Y YpreT flat; val eval_core' = substT Y ssig_T eval_core; val x' = substT Y Ypre_ssig_ssig_T x; val fst' = fst_const YpreT; val goal = mk_Trueprop_eq (eval_core $ (flat' $ x'), dead_pre_map' $ flat $ (eval_core' $ (dead_ssig_map' $ mk_convol (dead_ssig_map'' $ fst', eval_core) $ x'))); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_flat_tac ctxt ssig_induct' dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def = (trans OF [iffD2 OF [dtor_inject, HOLogic.mk_obj_eq eval_def RS fun_cong], dtor_unfold_thm]) |> unfold_thms ctxt [o_apply, eval_def RS symmetric_thm]; fun derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_ssig_T = Tsubst Y fp_ssig_T ssig_T; val dead_ssig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_ssig_map; val flat' = substT Y fpT flat; val x' = substT Y fp_ssig_ssig_T x; val goal = mk_Trueprop_eq (eval $ (flat' $ x'), eval $ (dead_ssig_map' $ eval $ x')); val cond_eval_o_flat = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (HOLogic.mk_comp (eval, flat')))] (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] RS fun_cong) OF [ext, ext]; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_flat_tac ctxt dead_pre_map_comp0 ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm cond_eval_o_flat)) |> Thm.close_derivation \<^here> end; fun derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_sig_T = Tsubst Y fp_ssig_T sig_T; val dead_sig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_sig_map; val Oper' = substT Y fpT Oper; val x' = substT Y fp_ssig_sig_T x; val goal = mk_Trueprop_eq (eval $ (Oper' $ x'), algLam $ (dead_sig_map' $ eval $ x')); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_Oper_tac ctxt live sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_eval_V_or_CLeaf ctxt Y fpT V_or_CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm = let val V_or_CLeaf' = substT Y fpT V_or_CLeaf; val x' = substT Y fpT x; val goal = mk_Trueprop_eq (eval $ (V_or_CLeaf' $ x'), x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_V_or_CLeaf_tac ctxt dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def = let val ssig_preT = Tsubst Y ssig_T preT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)] dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, cutSsig $ g'), HOLogic.mk_comp (dtor, extdd_f)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_mor_tac ctxt dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def prem)) |> Thm.close_derivation \<^here> end; fun derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val dead_ssig_map' = substYZ dead_ssig_map; val f' = substYZ f; val g' = substT Z ssig_preT g; val cutSsig_g = cutSsig $ g'; val id' = HOLogic.id_const ssig_T; val convol' = mk_convol (id', cutSsig_g); val dead_ssig_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, range_type (fastype_of convol'))] dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val eval_core_o_map = HOLogic.mk_comp (eval_core', dead_ssig_map'' $ convol'); val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ f', cutSsig_g), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, dead_ssig_map' $ f'), HOLogic.mk_comp (f', flat)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_mor_cutSsig_flat_tac ctxt eval_core_o_map dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm prem)) |> Thm.close_derivation \<^here> end; fun derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (extdd_f, VLeaf), f'); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_o_VLeaf_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf prem)) |> Thm.close_derivation \<^here> end; fun derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val g' = substT Z ssig_preT g; val corecU_g = corecU $ g'; val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ (extdd $ corecU_g), g'), HOLogic.mk_comp (dtor, corecU_g)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecU_pointfree_tac ctxt dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def)) |> Thm.close_derivation \<^here> end; fun derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def = let val corecU_pointfree = derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def; val corecU_thm = corecU_pointfree RS @{thm comp_eq_dest}; val corecU_ctor = let val arg_cong' = infer_instantiate' ctxt [NONE, NONE, SOME (Thm.cterm_of ctxt ctor)] arg_cong; in unfold_thms ctxt [ctor_dtor] (corecU_thm RS arg_cong') end; val corecU_unique = let val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val f' = substYZ f; val abs_f_o_VLeaf = Term.lambda f' (HOLogic.mk_comp (f', VLeaf)); val inject_refine' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt abs_f_o_VLeaf), SOME (Thm.cterm_of ctxt extdd)] @{thm inject_refine}; in unfold_thms ctxt @{thms atomize_imp} (((inject_refine' OF [extdd_o_VLeaf, extdd_o_VLeaf] RS iffD1) OF [asm_rl, corecU_pointfree]) OF [asm_rl, trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] OF [extdd_mor, corecU_pointfree RS extdd_mor]]) RS @{thm obj_distinct_prems} end; in (corecU_ctor, corecU_unique) end; fun derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_sig_T = Tsubst Y fpT sig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_sig_map; val Lam' = substT Y fpT Lam; val x' = substT Y fp_sig_T x; val goal = mk_Trueprop_eq (dtor $ (algLam $ x'), dead_pre_map' $ eval $ (Lam' $ (dead_sig_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algLam_tac ctxt pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_base ctxt Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def = let val fp_preT = Tsubst Y fpT preT; val proto_sctr' = substT Y fpT proto_sctr; val dead_pre_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_preT)] dead_pre_map; val dead_pre_map_dtor = dead_pre_map' $ dtor; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_base_tac ctxt dead_pre_map_dtor dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps = let val old_ssig_old_ssig_T = Tsubst Y old_ssig_T old_ssig_T; val dead_old_ssig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_ssig_map; val embL' = substT Y ssig_T embL; val x' = substT Y old_ssig_old_ssig_T x; val goal = mk_Trueprop_eq (flat $ (embL' $ (dead_old_ssig_map' $ embL $ x')), embL $ (old_flat $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_embL_tac ctxt old_ssig_induct' fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_ssig_T = Tsubst Y YpreT old_ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val embL' = substT Y YpreT embL; val x' = substT Y Ypre_old_ssig_T x; val goal = mk_Trueprop_eq (eval_core $ (embL' $ x'), dead_pre_map' $ embL $ (old_eval_core $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_embL_tac ctxt old_ssig_induct' dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL old_eval_core_simps eval_core_simps embL_simps embL_pointful_natural)) |> Thm.close_derivation \<^here> end; fun derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm = let val embL' = substT Y fpT embL; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, embL'), old_eval); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_embL_tac ctxt dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inx' = Inx_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inx'], old_algLam); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algLam_tac ctxt dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_k_as_ssig ctxt Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_k_T = Tsubst Y YpreT k_T; val k_as_ssig' = substT Y YpreT k_as_ssig; val x' = substT Y Ypre_k_T x; val goal = mk_Trueprop_eq (eval_core $ (k_as_ssig' $ x'), rho $ x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_k_as_ssig_tac ctxt pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algrho ctxt Y fpT Sig algLam algrho algLam_def algrho_def = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inr' = Inr_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inr'], algrho); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algrho_tac ctxt algLam_def algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_dtor_algrho ctxt Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val fppreT = Tsubst Y fpT YpreT; val fp_k_T = Tsubst Y fpT k_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_k_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_k_map; val rho' = substT Y fpT rho; val x' = substT Y fp_k_T x; val goal = mk_Trueprop_eq (dtor $ (algrho $ x'), dead_pre_map' $ eval $ (rho' $ (dead_k_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algrho_tac ctxt eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_step_or_merge ctxt Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam = let val proto_sctr' = substT Y fpT proto_sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); val algLam_algLam_pointful = mk_pointful ctxt algLam_algLam; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_step_tac ctxt proto_sctr_def old_algLam_pointful algLam_algLam_pointful)) |> Thm.close_derivation \<^here> end; fun derive_eval_sctr ctxt Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val sctr' = substT Y fpT sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, sctr'), HOLogic.mk_comp (ctor, dead_pre_map' $ eval)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_sctr_tac ctxt proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_pointfree_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, fp_ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, fpT)] dead_ssig_map; val dead_ssig_map'' = substT Z fpT dead_ssig_map; val f' = substT Z ssig_pre_ssig_T f; val g' = substT Z fpT g; val corecUU_f = corecUU $ f'; fun mk_eq fpf = mk_Trueprop_eq (fpf, Library.foldl1 HOLogic.mk_comp [eval, dead_ssig_map' $ Library.foldl1 HOLogic.mk_comp [ctor, dead_pre_map' $ eval, dead_pre_map'' $ (dead_ssig_map'' $ fpf)], f']); val corecUU_pointfree = let val goal = mk_eq corecUU_f; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecUU_pointfree_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor sctr_pointful_natural eval_sctr_pointful corecUU_def)) |> Thm.close_derivation \<^here> end; val corecUU_unique = let val prem = mk_eq g'; val goal = mk_Trueprop_eq (g', corecUU_f); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_corecUU_unique_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def prem)) |> Thm.close_derivation \<^here> end; in (corecUU_pointfree, corecUU_unique) end; fun define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer lthy = let val (flat_data as (flat, flat_def, _), lthy) = lthy |> define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map; val (eval_core_data as (eval_core, eval_core_def, _), lthy) = lthy |> define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam; val ((eval_data as (eval, _), cutSsig_data as (cutSsig, _)), lthy) = lthy |> define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core ||>> define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core; val ((algLam_data, unfold_data), lthy) = lthy |> define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval ||>> define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig; val flat_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R flat [flat_def] [] [sig_map_transfer]; val eval_core_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel ssig_rel ssig_rel eval_core eval_core_def fp_k_T_rel_eqs [pre_map_transfer, sig_map_transfer, ssig_map_transfer, flat_transfer, Lam_transfer, dtor_transfer]; in (((((((flat_data, flat_transfer), (eval_core_data, eval_core_transfer)), eval_data), cutSsig_data), algLam_data), unfold_data), lthy) end; fun derive_Sig_natural_etc ctxt live live_AsBs Y Z preT fpT k_or_fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf = let val SOME prod_bnf = bnf_of ctxt \<^type_name>\prod\; val f' = substT Z fpT f; val dead_ssig_map' = substT Z fpT dead_ssig_map; val extdd = Term.lambda f' (HOLogic.mk_comp (eval, dead_ssig_map' $ f')); val live_pre_map_def = map_def_of_bnf live_pre_bnf; val pre_map_comp = map_comp_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val dead_pre_map_cong = map_cong_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val sig_map_ident = map_ident_of_bnf sig_bnf; val sig_map_comp0 = map_comp0_of_bnf sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_cong = map_cong_of_bnf sig_bnf; val ssig_map_id = map_id_of_bnf ssig_bnf; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val dead_ssig_map_comp0 = map_comp0_of_bnf dead_ssig_bnf; val k_preT_map_id0s = map map_id0_of_bnf (map_filter (bnf_of ctxt) (fold add_type_namesT [preT, k_or_fpT] [])); val Sig_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Sig ([sig_map_thm, live_pre_map_def, @{thm BNF_Composition.id_bnf_def}] @ k_preT_map_id0s); val Oper_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Oper [Oper_map_thm]; val VLeaf_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f VLeaf [VLeaf_map_thm]; val Lam_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f Lam Lam_transfer [prod_bnf, pre_bnf, sig_bnf, ssig_bnf] []; val flat_natural = derive_natural_from_transfer ctxt live_AsBs [] fs f flat flat_transfer [ssig_bnf] []; val eval_core_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f eval_core eval_core_transfer [prod_bnf, pre_bnf, ssig_bnf] []; val Sig_pointful_natural = mk_pointful ctxt Sig_natural RS sym; val Oper_natural_pointful = mk_pointful ctxt Oper_natural; val Oper_pointful_natural = Oper_natural_pointful RS sym; val flat_pointful_natural = mk_pointful ctxt flat_natural RS sym; val Lam_natural_pointful = mk_pointful ctxt Lam_natural; val Lam_pointful_natural = Lam_natural_pointful RS sym; val eval_core_pointful_natural = mk_pointful ctxt eval_core_natural RS sym; val cutSsig_def_pointful_natural = mk_pointful ctxt (HOLogic.mk_obj_eq cutSsig_def) RS sym; val flat_VLeaf = derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps; val flat_flat = derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps; val eval_core_flat = derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps; val eval_thm = derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def; val eval_flat = derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm; val eval_Oper = derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def; val eval_VLeaf = derive_eval_V_or_CLeaf ctxt Y fpT VLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique VLeaf_map_thm eval_core_simps eval_thm; val eval_CLeaf = derive_eval_V_or_CLeaf ctxt Y fpT CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique CLeaf_map_thm eval_core_simps eval_thm; val extdd_mor = derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def; val mor_cutSsig_flat = derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm; val extdd_o_VLeaf = derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf; val (corecU_ctor, corecU_unique) = derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def; val dtor_algLam = derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def; in (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, [eval_Oper, eval_VLeaf, eval_CLeaf], corecU_ctor, corecU_unique, dtor_algLam) end; fun derive_embL_natural_etc ctxt Inx_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm = let val embL_natural = derive_natural_from_transfer ctxt [(Y, Z)] [] [] f embL embL_transfer [old_ssig_bnf, ssig_bnf] []; val embL_pointful_natural = mk_pointful ctxt embL_natural RS sym; val old_algLam_pointful = mk_pointful ctxt old_algLam_thm; val flat_embL = derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps; val eval_core_embL = derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps; val eval_embL = derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm; val algLam_algLam = derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam; in (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) end; fun define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar lthy = let val ssig_bnf = #fp_bnf ssig_fp_sugar; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val sctr_rhs = HOLogic.mk_comp (Oper, substT Y ssig_T proto_sctr); val ((sctr, sctr_def), lthy) = lthy |> define_const true fp_b version sctrN sctr_rhs; val (corecUU_data as (corecUU, corecUU_def), lthy) = lthy |> define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU; val eval_sctr = derive_eval_sctr lthy Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def; val sctr_transfer = derive_sctr_transfer lthy live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs [proto_sctr_transfer]; val sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f sctr sctr_transfer [pre_bnf, ssig_bnf] []; val sctr_pointful_natural = mk_pointful lthy sctr_natural RS sym; val eval_sctr_pointful = mk_pointful lthy eval_sctr RS sym; val (corecUU_pointfree, corecUU_unique) = derive_corecUU_pointfree_unique lthy Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def; val corecUU_thm = mk_pointful lthy corecUU_pointfree; val corecUU_transfer = derive_corecUU_transfer lthy live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs [pre_map_transfer, dtor_unfold_transfer, dtor_transfer, ssig_map_transfer, flat_transfer, eval_core_transfer, sctr_transfer, @{thm convol_transfer} (*FIXME: needed?*)]; in ((corecUU_data, corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) end; fun mk_equivp T = Const (\<^const_name>\equivp\, mk_predT [mk_pred2T T T]); fun derive_equivp_Retr ctxt fpT Retr R dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm = let val prem = HOLogic.mk_Trueprop (mk_equivp fpT $ R); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_equivp fpT $ (betapply (Retr, R)))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_equivp_Retr_tac ctxt dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm)) |> Thm.close_derivation \<^here> end; fun derive_Retr_coinduct ctxt fpT Retr R dtor_rel_coinduct_thm rel_eq_thm = let val goal = HOLogic.mk_Trueprop (list_all_free [R] (HOLogic.mk_imp (mk_leq R (Retr $ R), mk_leq R (HOLogic.eq_const fpT)))); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_Retr_coinduct_tac ctxt dtor_rel_coinduct_thm rel_eq_thm) |> Thm.close_derivation \<^here> end; fun derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy lthy = let val (R, _) = names_lthy |> yield_singleton (mk_Frees "R") (mk_pred2T fpT fpT); val pre_fpT = pre_type_of_ctor fpT ctor; val fp_pre_rel = mk_rel1 lthy fpT fpT pre_fpT dead_pre_bnf; val Retr = Abs ("R", fastype_of R, Abs ("a", fpT, Abs ("b", fpT, list_comb (fp_pre_rel, [Bound 2, dtor $ Bound 1, dtor $ Bound 0])))); val equivp_Retr = derive_equivp_Retr lthy fpT Retr R (rel_refl_of_bnf dead_pre_bnf) (rel_flip_of_bnf dead_pre_bnf) (rel_mono_of_bnf dead_pre_bnf) (rel_OO_of_bnf dead_pre_bnf); val Retr_coinduct = derive_Retr_coinduct lthy fpT Retr R (fp_sugar |> #fp_res |> #xtor_rel_co_induct) (fp_sugar |> #fp_bnf |> rel_eq_of_bnf); in (Retr, equivp_Retr, Retr_coinduct) end; fun mk_gen_cong fpT eval_domT = let val fp_relT = mk_pred2T fpT fpT in Const (\<^const_name>\cong.gen_cong\, [mk_predT [fp_relT, eval_domT, eval_domT], eval_domT --> fpT, fp_relT] ---> fp_relT) end; fun mk_cong_locale rel eval Retr = Const (\<^const_name>\cong\, mk_predT (map fastype_of [rel, eval, Retr])); fun derive_cong_locale ctxt rel eval Retr0 tac = let val Retr = enforce_type ctxt domain_type (domain_type (fastype_of rel)) Retr0; val goal = HOLogic.mk_Trueprop (list_comb (mk_cong_locale rel eval Retr, [rel, eval, Retr])); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => tac ctxt)) |> Thm.close_derivation \<^here> end; fun derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy = let val eval_domT = domain_type (fastype_of eval); fun cong_locale_tac ctxt = mk_cong_locale_tac ctxt (rel_mono_of_bnf dead_pre_bnf) (rel_map_of_bnf dead_pre_bnf) equivp_Retr (rel_mono_of_bnf dead_ssig_bnf) (rel_map_of_bnf dead_ssig_bnf) eval_thm eval_core_transfer; val rel = mk_rel1 lthy fpT fpT eval_domT dead_ssig_bnf; val cong_rhs = list_comb (mk_gen_cong fpT eval_domT, [rel, eval]); val ((_, cong_def), lthy) = lthy |> define_const false fp_b version congN cong_rhs; val cong_locale = derive_cong_locale lthy rel eval Retr cong_locale_tac; val fold_cong_def = Local_Defs.fold lthy [cong_def]; fun instance_of_gen thm = fold_cong_def (thm OF [cong_locale]); val cong_base = instance_of_gen @{thm cong.imp_gen_cong}; val cong_refl = instance_of_gen @{thm cong.gen_cong_reflp}; val cong_sym = instance_of_gen @{thm cong.gen_cong_symp}; val cong_trans = instance_of_gen @{thm cong.gen_cong_transp}; fun mk_cong_rho thm = thm RS instance_of_gen @{thm cong.gen_cong_rho}; val dtor_coinduct = @{thm predicate2I_obj} RS (Retr_coinduct RS instance_of_gen @{thm cong.coinduction} RS @{thm predicate2D_obj}); in (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) end; fun derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct lthy = let val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val dead_pre_map_id0 = map_id0_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_cong0 = map_cong0_of_bnf dead_pre_bnf; val dead_pre_map_cong0' = @{thm box_equals[OF _ o_apply[symmetric] id_apply[symmetric]]} RS dead_pre_map_cong0 RS ext; val dead_pre_rel_map = rel_map_of_bnf dead_pre_bnf; val ctor_alt_thm = eval_VLeaf RS (@{thm eq_comp_compI} OF [eval_sctr, trans OF [dead_pre_map_comp0 RS sym, trans OF [dead_pre_map_cong0', dead_pre_map_id0]]]); val cong_ctor_intro = mk_cong_rho ctor_alt_thm |> unfold_thms lthy [o_apply] |> (fn thm => sctr_transfer RS rel_funD RS thm) |> unfold_thms lthy (id_apply :: dead_pre_rel_map @ #rel_injects ssig_fp_bnf_sugar); in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = [cong_ctor_intro]}, lthy) end; fun update_cong_alg_intros ctxt cong_def cong_locale old_cong_def old_cong_locale emb = let fun instance_of_gen thm = Local_Defs.fold ctxt [cong_def] (thm OF [cong_locale]); fun instance_of_old_gen thm = Local_Defs.fold ctxt [old_cong_def] (thm OF [old_cong_locale]); val emb_idem = @{thm ord_le_eq_trans} OF [emb, instance_of_gen @{thm cong.gen_cong_idem}]; fun mk_rel_mono bnf = instance_of_old_gen @{thm cong.leq_gen_cong} RS rel_mono_of_bnf bnf RS @{thm predicate2D}; fun mk_intro bnf thm = mk_rel_mono bnf RS (@{thm predicate2D} OF [emb_idem, thm]); in map2 mk_intro end fun derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs lthy = let val old_cong_def = #cong_def old_dtor_coinduct_info; val old_cong_locale = #cong_locale old_dtor_coinduct_info; val old_cong_alg_intros = #cong_alg_intros old_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val cong_alg_intro = k_as_ssig_transfer RS rel_funD RS mk_cong_rho (HOLogic.mk_obj_eq algrho_def); val gen_cong_emb = (@{thm gen_cong_emb} OF [old_cong_locale, cong_locale, eval_embL, embL_transfer]) |> Local_Defs.fold lthy [old_cong_def, cong_def]; val cong_alg_intros = update_cong_alg_intros lthy cong_def cong_locale old_cong_def old_cong_locale gen_cong_emb old_all_dead_k_bnfs old_cong_alg_intros; in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_alg_intro :: cong_alg_intros}, lthy) end; fun derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs lthy = let val old1_cong_def = #cong_def old1_dtor_coinduct_info; val old1_cong_locale = #cong_locale old1_dtor_coinduct_info; val old1_cong_alg_intros = #cong_alg_intros old1_dtor_coinduct_info; val old2_cong_def = #cong_def old2_dtor_coinduct_info; val old2_cong_locale = #cong_locale old2_dtor_coinduct_info; val old2_cong_alg_intros = #cong_alg_intros old2_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, _, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val emb1 = (@{thm gen_cong_emb} OF [old1_cong_locale, cong_locale, eval_embLL, embLL_transfer]) |> Local_Defs.fold lthy [old1_cong_def, cong_def]; val emb2 = (@{thm gen_cong_emb} OF [old2_cong_locale, cong_locale, eval_embLR, embLR_transfer]) |> Local_Defs.fold lthy [old2_cong_def, cong_def]; val cong_alg_intros1 = update_cong_alg_intros lthy cong_def cong_locale old1_cong_def old1_cong_locale emb1 old1_all_dead_k_bnfs old1_cong_alg_intros; val cong_alg_intros2 = update_cong_alg_intros lthy cong_def cong_locale old2_cong_def old2_cong_locale emb2 old2_all_dead_k_bnfs old2_cong_alg_intros; val (cong_algrho_intros1, cong_ctor_intro1) = split_last cong_alg_intros1; val (cong_algrho_intros2, _) = split_last cong_alg_intros2; val (old1_all_rho_k_bnfs, old1_Sig_bnf) = split_last old1_all_dead_k_bnfs; val (old2_all_rho_k_bnfs, _) = split_last old2_all_dead_k_bnfs; val (friend_names, (cong_algrho_intros, all_rho_k_bnfs)) = merge_lists (op = o apply2 fst) (old1_friend_names ~~ (cong_algrho_intros1 ~~ old1_all_rho_k_bnfs)) (old2_friend_names ~~ (cong_algrho_intros2 ~~ old2_all_rho_k_bnfs)) |> split_list ||> split_list; in (({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_algrho_intros @ [cong_ctor_intro1]}, all_rho_k_bnfs @ [old1_Sig_bnf], friend_names), lthy) end; fun derive_corecUU_base fpT_name lthy = let val fp_sugar as {T = Type (_, fpT_args0), pre_bnf, fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val fp_alives = liveness_of_fp_bnf (length fpT_args0) fp_bnf; val (((Es, Fs0), [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; val Fs = @{map 3} (fn alive => fn E => fn F => if alive then F else E) fp_alives Es Fs0; val As = Es @ [Y]; val Bs = Es @ [Z]; val live_EsFs = filter (op <>) (Es ~~ Fs); val live_AsBs = live_EsFs @ [(Y, Z)]; val fTs = map (op -->) live_EsFs; val RTs = map (uncurry mk_pred2T) live_EsFs; val live = length live_EsFs; val ((((((x, fs), f), g), Rs), R), names_lthy) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> mk_Frees "f" fTs ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> mk_Frees "R" RTs ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val ctor = mk_ctor Es (the_single (#ctors fp_res)); val dtor = mk_dtor Es (the_single (#dtors fp_res)); val fpT = Type (fpT_name, Es); val preT = pre_type_of_ctor Y ctor; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Es Y preT ||>> define_ssig_type fp_b version fp_alives Es Y fpT; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (((dead_pre_bnf, dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 pre_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val sig_T = Type (sig_T_name, As); val ssig_T = Type (ssig_T_name, As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_sig_map = mk_map 1 As Bs (map_of_bnf dead_sig_bnf); val [Oper, VLeaf, CLeaf] = map (mk_ctr As) (#ctrs ssig_ctr_sugar); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 As Bs (map_of_bnf dead_ssig_bnf); val ((Lam, Lam_def), lthy) = lthy |> define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf; val proto_sctr = Sig; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val pre_rel_def = rel_def_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val [dtor_rel_thm] = #xtor_rels fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val dtor_transfer = derive_dtor_transfer lthy live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm; val preT_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (add_type_namesT preT [])); val Sig_transfer = derive_sig_transfer I lthy live_AsBs pre_rel sig_rel Rs R Sig pre_rel_def preT_rel_eqs (the_single (#ctr_transfers sig_fp_ctr_sugar)); val proto_sctr_transfer = Sig_transfer; val unsig_transfer = derive_sig_transfer swap lthy live_AsBs pre_rel sig_rel Rs R unsig pre_rel_def preT_rel_eqs (the_single (#sel_transfers sig_fp_ctr_sugar)); val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, sig_map_transfer, Sig_transfer, unsig_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, eval_VLeaf, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_pointful_natural = Sig_pointful_natural; val algLam_thm = derive_algLam_base lthy Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (Retr, equivp_Retr, Retr_coinduct) = lthy |> derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = Sig, friends = Symtab.empty}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (dtor_transferN, [dtor_transfer]), (equivp_RetrN, [equivp_Retr]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (Retr_coinductN, [Retr_coinduct]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = [], sig_fp_sugars = [sig_fp_sugar], ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = [algLam_thm], algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = [dead_pre_bnf], Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_step (fpT as Type (fpT_name, res_Ds)) ({friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugars as old_sig_fp_sugar :: _, ssig_fp_sugar = old_ssig_fp_sugar, Lam = old_Lam0, proto_sctr = old_proto_sctr0, flat = old_flat0, eval_core = old_eval_core0, eval = old_eval0, algLam = old_algLam0, dtor_transfer, Lam_transfer = old_Lam_transfer, Lam_pointful_natural = old_Lam_pointful_natural, proto_sctr_transfer = old_proto_sctr_transfer, flat_simps = old_flat_simps, eval_core_simps = old_eval_core_simps, eval_thm = old_eval_thm, all_algLam_algs = old_all_algLam_algs, algLam_thm = old_algLam_thm, dtor_algLam = old_dtor_algLam, buffer = old_buffer, all_dead_k_bnfs = old_all_dead_k_bnfs, Retr = old_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old_dtor_coinduct_info, ...} : corec_info) friend_name friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy = let val {pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val names_lthy = lthy |> fold Variable.declare_typ [Y, Z]; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val friend_names = friend_name :: old_friend_names; val old_sig_bnf = #fp_bnf old_sig_fp_sugar; val old_ssig_bnf = #fp_bnf old_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((dead_pre_bnf, dead_fp_bnf), dead_old_sig_bnf), dead_old_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old_sig_bnf ||>> bnf_kill_all_but 1 old_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old_ssig_fp_ctr_sugar = #fp_ctr_sugar old_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old_ssig_fp_bnf_sugar = #fp_bnf_sugar old_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old_ssig_ctr_sugar = #ctr_sugar old_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_ssig_T_name = fst (dest_Type (#T old_ssig_fp_sugar)); val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val old_sig_T = Type (old_sig_T_name, res_As); val old_ssig_T = Type (old_ssig_T_name, res_As); val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val old_Lam_domT = Tsubst Y YpreT old_sig_T; val old_eval_core_domT = Tsubst Y YpreT old_ssig_T; val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val dead_k_map = mk_map1 lthy Y Z k_T dead_k_bnf; val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old_Oper, old_VLeaf, old_CLeaf] = map (mk_ctr res_As) (#ctrs old_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val dead_old_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old_Lam = enforce_type lthy domain_type old_Lam_domT old_Lam0; val old_proto_sctr = enforce_type lthy domain_type preT old_proto_sctr0; val old_flat = enforce_type lthy range_type old_ssig_T old_flat0; val old_eval_core = enforce_type lthy domain_type old_eval_core_domT old_eval_core0; val old_eval = enforce_type lthy range_type fpT old_eval0; val old_algLam = enforce_type lthy range_type fpT old_algLam0; val ((embL, embL_def, embL_simps), lthy) = lthy |> define_embL embLN fp_b version Y Z fpT old_sig_T old_ssig_T k_T ssig_T Inl_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old_sig_T k_T Sig old_proto_sctr; val pre_map_comp = map_comp_of_bnf pre_bnf; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val fp_k_T_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (fold add_type_namesT [fpT, k_T] [])); val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old_sig_map_comp = map_comp_of_bnf old_sig_bnf; val old_sig_map_cong = map_cong_of_bnf old_sig_bnf; val old_ssig_map_thms = #map_thms old_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old_sig_map_transfer = map_transfer_of_bnf old_sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old_ssig_induct = the_single (#co_inducts old_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs [old_proto_sctr_transfer]; val embL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embL [embL_def] fp_k_T_rel_eqs [old_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def fp_k_T_rel_eqs [pre_map_transfer, old_Lam_transfer, embL_transfer, rho_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT k_T sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) = derive_embL_natural_etc lthy Inl_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam; val k_as_ssig = mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf; val k_as_ssig' = substT Y fpT k_as_ssig; val algrho_rhs = HOLogic.mk_comp (eval, k_as_ssig'); val ((algrho, algrho_def), lthy) = lthy |> define_const true fp_b version algrhoN algrho_rhs; val k_as_ssig_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R k_as_ssig [] fp_k_T_rel_eqs [sig_map_transfer]; val k_as_ssig_natural = derive_natural_from_transfer lthy [(Y, Z)] [] [] f k_as_ssig k_as_ssig_transfer [ssig_bnf] [dead_k_bnf]; val k_as_ssig_natural_pointful = mk_pointful lthy k_as_ssig_natural; val [_, Lam_Inr] = derive_Lam_Inl_Inr lthy Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def; val eval_core_k_as_ssig = derive_eval_core_k_as_ssig lthy Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps; val algLam_algrho = derive_algLam_algrho lthy Y fpT Sig algLam algrho algLam_def algrho_def; val dtor_algrho = derive_dtor_algrho lthy Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def; val all_algLam_algs = algLam_algLam :: algLam_algrho :: old_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val Retr = enforce_type lthy (domain_type o domain_type) fpT old_Retr0; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLam]), (algLam_algrhoN, [algLam_algrho]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_algrhoN, [dtor_algrho]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (embL_pointful_naturalN, [embL_pointful_natural]), (embL_transferN, [embL_transfer]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (k_as_ssig_naturalN, [k_as_ssig_natural]), (k_as_ssig_transferN, [k_as_ssig_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (rho_transferN, [rho_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); val phi = Local_Theory.target_morphism lthy; in (({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_sig_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = dead_k_bnf :: old_all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info phi, ({algrho = algrho, dtor_algrho = dtor_algrho, algLam_algrho = algLam_algrho} |> morph_friend_info phi)), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_merge (fpT as Type (fpT_name, res_Ds)) ({friend_names = old1_friend_names, sig_fp_sugars = old1_sig_fp_sugars as old1_sig_fp_sugar :: _, ssig_fp_sugar = old1_ssig_fp_sugar, Lam = old1_Lam0, proto_sctr = old1_proto_sctr0, flat = old1_flat0, eval_core = old1_eval_core0, eval = old1_eval0, algLam = old1_algLam0, dtor_transfer, Lam_transfer = old1_Lam_transfer, Lam_pointful_natural = old1_Lam_pointful_natural, proto_sctr_transfer = old1_proto_sctr_transfer, flat_simps = old1_flat_simps, eval_core_simps = old1_eval_core_simps, eval_thm = old1_eval_thm, all_algLam_algs = old1_all_algLam_algs, algLam_thm = old1_algLam_thm, dtor_algLam = old1_dtor_algLam, buffer = old1_buffer, all_dead_k_bnfs = old1_all_dead_k_bnfs, Retr = old1_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old1_dtor_coinduct_info, ...} : corec_info) ({friend_names = old2_friend_names, sig_fp_sugars = old2_sig_fp_sugars as old2_sig_fp_sugar :: _, ssig_fp_sugar = old2_ssig_fp_sugar, Lam = old2_Lam0, flat = old2_flat0, eval_core = old2_eval_core0, eval = old2_eval0, algLam = old2_algLam0, Lam_transfer = old2_Lam_transfer, Lam_pointful_natural = old2_Lam_pointful_natural, flat_simps = old2_flat_simps, eval_core_simps = old2_eval_core_simps, eval_thm = old2_eval_thm, all_algLam_algs = old2_all_algLam_algs, algLam_thm = old2_algLam_thm, dtor_algLam = old2_dtor_algLam, buffer = old2_buffer, all_dead_k_bnfs = old2_all_dead_k_bnfs, dtor_coinduct_info = old2_dtor_coinduct_info, ...} : corec_info) lthy = let val {T = Type (_, fpT_args0), pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf (length fpT_args0) live_fp_bnf; val ((Ds, [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val old1_sig_T_name = fst (dest_Type (#T old1_sig_fp_sugar)); val old2_sig_T_name = fst (dest_Type (#T old2_sig_fp_sugar)); val old1_ssig_T_name = fst (dest_Type (#T old1_ssig_fp_sugar)); val old2_ssig_T_name = fst (dest_Type (#T old2_ssig_fp_sugar)); val fp_alives = map (K false) live_fp_alives; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val fpT0 = Type (fpT_name, Ds); val old1_sig_T0 = Type (old1_sig_T_name, As); val old2_sig_T0 = Type (old2_sig_T_name, As); val old1_sig_T = Type (old1_sig_T_name, res_As); val old2_sig_T = Type (old2_sig_T_name, res_As); val old1_ssig_T = Type (old1_ssig_T_name, res_As); val old2_ssig_T = Type (old2_ssig_T_name, res_As); val old1_Lam_domT = Tsubst Y YpreT old1_sig_T; val old2_Lam_domT = Tsubst Y YpreT old2_sig_T; val old1_eval_core_domT = Tsubst Y YpreT old1_ssig_T; val old2_eval_core_domT = Tsubst Y YpreT old2_ssig_T; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old1_sig_T0, old2_sig_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val old1_sig_bnf = #fp_bnf old1_sig_fp_sugar; val old2_sig_bnf = #fp_bnf old2_sig_fp_sugar; val old1_ssig_bnf = #fp_bnf old1_ssig_fp_sugar; val old2_ssig_bnf = #fp_bnf old2_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((((dead_pre_bnf, dead_fp_bnf), dead_old1_sig_bnf), dead_old2_sig_bnf), dead_old1_ssig_bnf), dead_old2_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old1_sig_bnf ||>> bnf_kill_all_but 1 old2_sig_bnf ||>> bnf_kill_all_but 1 old1_ssig_bnf ||>> bnf_kill_all_but 1 old2_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old1_ssig_fp_ctr_sugar = #fp_ctr_sugar old1_ssig_fp_sugar; val old2_ssig_fp_ctr_sugar = #fp_ctr_sugar old2_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old1_ssig_fp_bnf_sugar = #fp_bnf_sugar old1_ssig_fp_sugar; val old2_ssig_fp_bnf_sugar = #fp_bnf_sugar old2_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old1_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old1_ssig_fp_sugar); val old2_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old2_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old1_ssig_ctr_sugar = #ctr_sugar old1_ssig_fp_ctr_sugar; val old2_ssig_ctr_sugar = #ctr_sugar old2_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old1_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_sig_bnf); val dead_old2_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old1_Oper, old1_VLeaf, old1_CLeaf] = map (mk_ctr res_As) (#ctrs old1_ssig_ctr_sugar); val [old2_Oper, old2_VLeaf, old2_CLeaf] = map (mk_ctr res_As) (#ctrs old2_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val old1_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_ssig_bnf); val old2_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old1_Lam = enforce_type lthy domain_type old1_Lam_domT old1_Lam0; val old2_Lam = enforce_type lthy domain_type old2_Lam_domT old2_Lam0; val old1_proto_sctr = enforce_type lthy domain_type preT old1_proto_sctr0; val old1_flat = enforce_type lthy range_type old1_ssig_T old1_flat0; val old2_flat = enforce_type lthy range_type old2_ssig_T old2_flat0; val old1_eval_core = enforce_type lthy domain_type old1_eval_core_domT old1_eval_core0; val old2_eval_core = enforce_type lthy domain_type old2_eval_core_domT old2_eval_core0; val old1_eval = enforce_type lthy range_type fpT old1_eval0; val old2_eval = enforce_type lthy range_type fpT old2_eval0; val old1_algLam = enforce_type lthy range_type fpT old1_algLam0; val old2_algLam = enforce_type lthy range_type fpT old2_algLam0; val ((embLL, embLL_def, embLL_simps), lthy) = lthy |> define_embL embLLN fp_b version Y Z fpT old1_sig_T old1_ssig_T old2_sig_T ssig_T Inl_const dead_old1_sig_map Sig old1_Oper old1_VLeaf old1_CLeaf Oper VLeaf CLeaf; val ((embLR, embLR_def, embLR_simps), lthy) = lthy |> define_embL embLRN fp_b version Y Z fpT old2_sig_T old2_ssig_T old1_sig_T ssig_T (fn T => fn U => Inr_const U T) dead_old2_sig_map Sig old2_Oper old2_VLeaf old2_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old1_sig_T old2_sig_T Sig old1_proto_sctr; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old1_sig_map_comp = map_comp_of_bnf old1_sig_bnf; val old2_sig_map_comp = map_comp_of_bnf old2_sig_bnf; val old1_sig_map_cong = map_cong_of_bnf old1_sig_bnf; val old2_sig_map_cong = map_cong_of_bnf old2_sig_bnf; val old1_ssig_map_thms = #map_thms old1_ssig_fp_bnf_sugar; val old2_ssig_map_thms = #map_thms old2_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old1_sig_map_transfer = map_transfer_of_bnf old1_sig_bnf; val old2_sig_map_transfer = map_transfer_of_bnf old2_sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old1_ssig_induct = the_single (#co_inducts old1_ssig_fp_induct_sugar); val old2_ssig_induct = the_single (#co_inducts old2_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def [] [old1_proto_sctr_transfer]; val embLL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLL [embLL_def] [] [old1_sig_map_transfer]; val embLR_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLR [embLR_def] [] [old2_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, old1_Lam_transfer, old2_Lam_transfer, embLL_transfer, embLR_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embLL_pointful_natural, old1_algLam_pointful, eval_embLL, algLam_algLamL) = derive_embL_natural_etc lthy Inl_const old1_ssig_bnf ssig_bnf Y Z preT fpT old1_ssig_T ssig_T dead_pre_map Sig old1_ssig_map embLL old1_algLam algLam old1_flat flat old1_eval_core eval_core old1_eval eval x f old1_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old1_sig_map_comp old1_sig_map_cong old1_ssig_map_thms old1_Lam_pointful_natural Lam_def old1_flat_simps flat_simps embLL_simps embLL_transfer old1_eval_core_simps eval_core_simps old1_eval_thm eval_thm old1_dtor_algLam dtor_algLam old1_algLam_thm; val (embLR_pointful_natural, _, eval_embLR, algLam_algLamR) = derive_embL_natural_etc lthy Inr_const old2_ssig_bnf ssig_bnf Y Z preT fpT old2_ssig_T ssig_T dead_pre_map Sig old2_ssig_map embLR old2_algLam algLam old2_flat flat old2_eval_core eval_core old2_eval eval x f old2_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old2_sig_map_comp old2_sig_map_cong old2_ssig_map_thms old2_Lam_pointful_natural Lam_def old2_flat_simps flat_simps embLR_simps embLR_transfer old2_eval_core_simps eval_core_simps old2_eval_thm eval_thm old2_dtor_algLam dtor_algLam old2_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old1_algLam_pointful algLam_algLamL; val all_algLam_algs = algLam_algLamL :: algLam_algLamR :: merge_lists (Thm.eq_thm_prop o apply2 zero_var_indexes) old1_all_algLam_algs old2_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val Retr = enforce_type lthy (domain_type o domain_type) fpT old1_Retr0; val embed_Sig_inl = embed_Sig lthy Sig (Inl_const old1_sig_T old2_sig_T); val embed_Sig_inr = embed_Sig lthy Sig (Inr_const old1_sig_T old2_sig_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old1_buffer); val friends = Symtab.merge (K true) (Symtab.map (K (apsnd embed_Sig_inl)) (#friends old1_buffer), Symtab.map (K (apsnd embed_Sig_inr)) (#friends old2_buffer)); val old_fp_sugars = merge_lists (op = o apply2 (fst o dest_Type o #T)) old1_sig_fp_sugars old2_sig_fp_sugars; val ((dtor_coinduct_info, all_dead_k_bnfs, friend_names), lthy) = lthy |> derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLamL, algLam_algLamR]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (embL_pointful_naturalN, [embLL_pointful_natural, embLR_pointful_natural]), (embL_transferN, [embLL_transfer, embLR_transfer]), (evalN, [eval_thm]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun set_corec_info_exprs fpT_name f = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi => let val exprs = f phi in Data.map (apsnd (fn [info_tab] => [Symtab.map_default (fpT_name, exprs) (K exprs) info_tab])) end); fun subsume_corec_info_ad ctxt {fpT = fpT1, friend_names = friend_names1} {fpT = fpT2, friend_names = friend_names2} = Sign.typ_instance (Proof_Context.theory_of ctxt) (fpT1, fpT2) andalso subset (op =) (friend_names1, friend_names2); fun subsume_corec_info_expr ctxt expr1 expr2 = subsume_corec_info_ad ctxt (corec_ad_of_expr expr1) (corec_ad_of_expr expr2); fun instantiate_corec_info thy res_T (info as {fpT, ...}) = let val As_rho = tvar_subst thy [fpT] [res_T]; val substAT = Term.typ_subst_TVars As_rho; val substA = Term.subst_TVars As_rho; val phi = Morphism.typ_morphism "BNF" substAT $> Morphism.term_morphism "BNF" substA; in morph_corec_info phi info end; fun instantiate_corec_info_expr thy res_T (Ad ({friend_names, ...}, f)) = Ad ({fpT = res_T, friend_names = friend_names}, f #>> instantiate_corec_info thy res_T) | instantiate_corec_info_expr thy res_T (Info info) = Info (instantiate_corec_info thy res_T info); fun ensure_Info expr = corec_info_of_expr expr #>> Info and ensure_Info_if_Info old_expr (expr, lthy) = if is_Info old_expr then ensure_Info expr lthy else (expr, lthy) and merge_corec_info_exprs old_exprs expr1 expr2 lthy = if subsume_corec_info_expr lthy expr2 expr1 then if subsume_corec_info_expr lthy expr1 expr2 andalso is_Ad expr1 then (expr2, lthy) else ensure_Info_if_Info expr2 (expr1, lthy) else if subsume_corec_info_expr lthy expr1 expr2 then ensure_Info_if_Info expr1 (expr2, lthy) else let val thy = Proof_Context.theory_of lthy; val {fpT = fpT1, friend_names = friend_names1} = corec_ad_of_expr expr1; val {fpT = fpT2, friend_names = friend_names2} = corec_ad_of_expr expr2; val fpT0 = typ_unify_disjointly thy (fpT1, fpT2); val fpT = singleton (freeze_types lthy []) fpT0; val friend_names = merge_lists (op =) friend_names1 friend_names2; val expr = Ad ({fpT = fpT, friend_names = friend_names}, corec_info_of_expr expr1 ##>> corec_info_of_expr expr2 #-> uncurry (derive_corecUU_merge fpT)); val old_same_type_exprs = if old_exprs then [] |> Sign.typ_instance thy (fpT1, fpT0) ? cons expr1 |> Sign.typ_instance thy (fpT2, fpT0) ? cons expr2 else []; in (expr, lthy) |> fold ensure_Info_if_Info old_same_type_exprs end and insert_corec_info_expr expr exprs lthy = let val thy = Proof_Context.theory_of lthy; val {fpT = new_fpT, ...} = corec_ad_of_expr expr; val is_Tinst = curry (Sign.typ_instance thy); fun is_Tequiv T U = is_Tinst T U andalso is_Tinst U T; val (((equiv_exprs, sub_exprs), sup_exprs), incomp_exprs) = exprs |> List.partition ((fn {fpT, ...} => is_Tequiv fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst new_fpT fpT) o corec_ad_of_expr); fun add_instantiated_incomp_expr expr exprs = let val {fpT, ...} = corec_ad_of_expr expr in (case try (typ_unify_disjointly thy) (fpT, new_fpT) of SOME new_T => let val subsumes = (fn {fpT, ...} => is_Tinst new_T fpT) o corec_ad_of_expr in if exists (exists subsumes) [exprs, sub_exprs] then exprs else instantiate_corec_info_expr thy new_T expr :: exprs end | NONE => exprs) end; val unincomp_exprs = fold add_instantiated_incomp_expr incomp_exprs []; val ((merged_sub_exprs, merged_unincomp_exprs), lthy) = lthy |> fold_map (merge_corec_info_exprs true expr) sub_exprs ||>> fold_map (merge_corec_info_exprs false expr) unincomp_exprs; val (merged_equiv_expr, lthy) = (expr, lthy) |> fold (uncurry o merge_corec_info_exprs true) equiv_exprs; in (merged_unincomp_exprs @ merged_sub_exprs @ merged_equiv_expr :: sup_exprs @ incomp_exprs |> sort (rev_order o int_ord o apply2 (length o #friend_names o corec_ad_of_expr)), lthy) end and register_corec_info (info as {fpT = Type (fpT_name, _), ...}) lthy = let val (exprs, lthy) = insert_corec_info_expr (Info info) (corec_info_exprs_of lthy fpT_name) lthy; in lthy |> set_corec_info_exprs fpT_name (fn phi => map (morph_corec_info_expr phi) exprs) end and corec_info_of_expr (Ad (_, f)) lthy = f lthy | corec_info_of_expr (Info info) lthy = (info, lthy); fun nonempty_corec_info_exprs_of fpT_name lthy = (case corec_info_exprs_of lthy fpT_name of [] => derive_corecUU_base fpT_name lthy |> (fn (info, lthy) => ([Info info], lthy |> set_corec_info_exprs fpT_name (fn phi => [Info (morph_corec_info phi info)]))) | exprs => (exprs, lthy)); fun corec_info_of res_T lthy = (case res_T of Type (fpT_name, _) => let val (exprs, lthy) = nonempty_corec_info_exprs_of fpT_name lthy; val thy = Proof_Context.theory_of lthy; val expr = (case find_first ((fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) o corec_ad_of_expr) exprs of SOME expr => expr | NONE => error ("Invalid type: " ^ Syntax.string_of_typ lthy res_T)); val (info, lthy) = corec_info_of_expr expr lthy; in (instantiate_corec_info thy res_T info, lthy |> is_Ad expr ? register_corec_info info) end | _ => not_codatatype lthy res_T); fun maybe_corec_info_of ctxt res_T = (case res_T of Type (fpT_name, _) => let val thy = Proof_Context.theory_of ctxt; val infos = corec_infos_of ctxt fpT_name; in find_first (fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) infos |> Option.map (instantiate_corec_info thy res_T) end | _ => not_codatatype ctxt res_T); fun prepare_friend_corec friend_name friend_T lthy = let val (arg_Ts, res_T) = strip_type friend_T; val Type (fpT_name, res_Ds) = (case res_T of T as Type _ => T | T => error (not_codatatype lthy T)); val _ = not (null arg_Ts) orelse error "Function with no argument cannot be registered as friend"; val {T = Type (fpT_name, fpT_args0), pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val num_fp_tyargs = length fpT_args0; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf num_fp_tyargs live_fp_bnf; val (old_info as {friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugar :: _, buffer = old_buffer, ...}, lthy) = corec_info_of res_T lthy; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_sig_alives = liveness_of_fp_bnf (num_fp_tyargs + 1) (#fp_bnf old_sig_fp_sugar); (* FIXME: later *) val fp_alives = fst (split_last old_sig_alives); val fp_alives = map (K false) live_fp_alives; val _ = not (member (op =) old_friend_names friend_name) orelse error ("Function " ^ quote (Syntax.string_of_term lthy (Const (friend_name, friend_T))) ^ " already registered as friend"); val lthy = lthy |> Variable.declare_typ friend_T; val ((Ds, [Y, Z]), _) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val dead_Ds = Ds; val live_As = [Y]; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val fpT0 = Type (fpT_name, Ds); val k_Ts0 = map (typ_subst_nonatomic (res_Ds ~~ Ds) o typ_subst_nonatomic [(res_T, Y)]) arg_Ts; val k_T0 = mk_tupleT_balanced k_Ts0; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val k_As = fold Term.add_tfreesT k_Ts0 []; val _ = (case filter_out (member (op =) As o TFree) k_As of [] => () | k_A :: _ => error ("Cannot have type variable " ^ quote (Syntax.string_of_typ lthy (TFree k_A)) ^ " in the argument types when it does not occur as an immediate argument of the result \ \type constructor")); val substDT = Term.typ_subst_atomic (Ds ~~ res_Ds); val old_sig_T0 = Type (old_sig_T_name, As); val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val (((dead_k_bnf, sig_fp_sugar), ssig_fp_sugar), lthy) = lthy |> bnf_with_deads_and_lives dead_Ds live_As Y fpT0 k_T0 ||>> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old_sig_T0, k_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val _ = live_of_bnf dead_k_bnf = 1 orelse error "Impossible type for friend (the result codatatype must occur live in the arguments)"; val (dead_pre_bnf, lthy) = lthy |> bnf_kill_all_but 1 pre_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val preT = pre_type_of_ctor Y ctor; val old_sig_T = substDT old_sig_T0; val k_T = substDT k_T0; val ssig_T = Type (ssig_T_name, res_As); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; in ((old_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy) end; fun register_friend_corec key fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar friend_const rho rho_transfer old_info lthy = let val friend_T = fastype_of friend_const; val res_T = body_type friend_T; in derive_corecUU_step res_T old_info key friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy |> (fn ((info, friend_info), lthy) => (friend_info, register_corec_info info lthy)) end; fun merge_corec_info_exprss exprs1 exprs2 lthy = let fun all_friend_names_of exprs = fold (union (op =)) (map (#friend_names o corec_ad_of_expr) exprs) []; val friend_names1 = all_friend_names_of exprs1; val friend_names2 = all_friend_names_of exprs2; in if subset (op =) (friend_names2, friend_names1) then if subset (op =) (friend_names1, friend_names2) andalso length (filter is_Info exprs2) > length (filter is_Info exprs1) then (exprs2, lthy) else (exprs1, lthy) else if subset (op =) (friend_names1, friend_names2) then (exprs2, lthy) else fold_rev (uncurry o insert_corec_info_expr) exprs2 (exprs1, lthy) end; fun merge_corec_info_tabs info_tab1 info_tab2 lthy = let val fpT_names = union (op =) (Symtab.keys info_tab1) (Symtab.keys info_tab2); fun add_infos_of fpT_name (info_tab, lthy) = (case Symtab.lookup info_tab1 fpT_name of NONE => (Symtab.update_new (fpT_name, the (Symtab.lookup info_tab2 fpT_name)) info_tab, lthy) | SOME exprs1 => (case Symtab.lookup info_tab2 fpT_name of NONE => (Symtab.update_new (fpT_name, exprs1) info_tab, lthy) | SOME exprs2 => let val (exprs, lthy) = merge_corec_info_exprss exprs1 exprs2 lthy in (Symtab.update_new (fpT_name, exprs) info_tab, lthy) end)); in fold add_infos_of fpT_names (Symtab.empty, lthy) end; fun consolidate lthy = (case snd (Data.get (Context.Proof lthy)) of [_] => raise Same.SAME | info_tab :: info_tabs => let val (info_tab', lthy) = fold_rev (uncurry o merge_corec_info_tabs) info_tabs (info_tab, lthy); in - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi => Data.map (apsnd (fn _ => [Symtab.map (K (map (morph_corec_info_expr phi))) info_tab']))) lthy end); fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML @@ -1,2391 +1,2394 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Author: Aymeric Bouzy, Ecole polytechnique Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor sugar ("corec" and friends). *) signature BNF_GFP_GREC_SUGAR = sig datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option val parse_corec_equation: Proof.context -> term list -> term -> term list * term val explore_corec_equation: Proof.context -> bool -> bool -> string -> term -> BNF_GFP_Grec_Sugar_Util.s_parse_info -> typ -> term list * term -> term list * term val build_corecUU_arg_and_goals: bool -> term -> term list * term -> local_theory -> (((thm list * thm list * thm list) * term list) * term) * local_theory val derive_eq_corecUU: Proof.context -> BNF_GFP_Grec.corec_info -> term -> term -> thm -> thm val derive_unique: Proof.context -> morphism -> term -> BNF_GFP_Grec.corec_info -> string -> thm -> thm val corec_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> local_theory val corecursive_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> Proof.state val friend_of_corec_cmd: (string * string option) * string -> local_theory -> Proof.state val coinduction_upto_cmd: string * string -> local_theory -> local_theory end; structure BNF_GFP_Grec_Sugar : BNF_GFP_GREC_SUGAR = struct open Ctr_Sugar open BNF_Util open BNF_Tactics open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_FP_N2M_Sugar open BNF_GFP_Util open BNF_GFP_Rec_Sugar open BNF_FP_Rec_Sugar_Transfer open BNF_GFP_Grec open BNF_GFP_Grec_Sugar_Util open BNF_GFP_Grec_Sugar_Tactics val cong_N = "cong_"; val baseN = "base"; val reflN = "refl"; val symN = "sym"; val transN = "trans"; val cong_introsN = prefix cong_N "intros"; val codeN = "code"; val coinductN = "coinduct"; val coinduct_uptoN = "coinduct_upto"; val corecN = "corec"; val ctrN = "ctr"; val discN = "disc"; val disc_iffN = "disc_iff"; val eq_algrhoN = "eq_algrho"; val eq_corecUUN = "eq_corecUU"; val friendN = "friend"; val inner_elimN = "inner_elim"; val inner_inductN = "inner_induct"; val inner_simpN = "inner_simp"; val rhoN = "rho"; val selN = "sel"; val uniqueN = "unique"; val inner_fp_suffix = "_inner_fp"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val unfold_id_thms1 = map (fn thm => thm RS eq_reflection) @{thms id_bnf_o o_id_bnf id_apply o_apply} @ @{thms fst_def[abs_def, symmetric] snd_def[abs_def, symmetric]}; fun unfold_id_bnf_etc lthy = let val thy = Proof_Context.theory_of lthy in Raw_Simplifier.rewrite_term thy unfold_id_thms1 [] #> Raw_Simplifier.rewrite_term thy @{thms BNF_Composition.id_bnf_def} [] end; datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option; val corec_option_parser = Parse.group (K "option") (Plugin_Name.parse_filter >> Plugins_Option || Parse.reserved "friend" >> K Friend_Option || Parse.reserved "transfer" >> K Transfer_Option); type codatatype_extra = {case_dtor: thm, case_trivial: thm, abs_rep_transfers: thm list}; fun morph_codatatype_extra phi ({case_dtor, case_trivial, abs_rep_transfers} : codatatype_extra) = {case_dtor = Morphism.thm phi case_dtor, case_trivial = Morphism.thm phi case_trivial, abs_rep_transfers = map (Morphism.thm phi) abs_rep_transfers}; val transfer_codatatype_extra = morph_codatatype_extra o Morphism.transfer_morphism; type coinduct_extra = {coinduct: thm, coinduct_attrs: Token.src list, cong_intro_pairs: (string * thm) list}; fun morph_coinduct_extra phi ({coinduct, coinduct_attrs, cong_intro_pairs} : coinduct_extra) = {coinduct = Morphism.thm phi coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = map (apsnd (Morphism.thm phi)) cong_intro_pairs}; val transfer_coinduct_extra = morph_coinduct_extra o Morphism.transfer_morphism; type friend_extra = {eq_algrhos: thm list, algrho_eqs: thm list}; val empty_friend_extra = {eq_algrhos = [], algrho_eqs = []}; fun merge_friend_extras ({eq_algrhos = eq_algrhos1, algrho_eqs = algrho_eqs1}, {eq_algrhos = eq_algrhos2, algrho_eqs = algrho_eqs2}) = {eq_algrhos = union Thm.eq_thm_prop eq_algrhos1 eq_algrhos2, algrho_eqs = union Thm.eq_thm_prop algrho_eqs1 algrho_eqs2}; type corec_sugar_data = codatatype_extra Symtab.table * coinduct_extra Symtab.table * friend_extra Symtab.table; structure Data = Generic_Data ( type T = corec_sugar_data; val empty = (Symtab.empty, Symtab.empty, Symtab.empty); fun merge data : T = (Symtab.merge (K true) (apply2 #1 data), Symtab.merge (K true) (apply2 #2 data), Symtab.join (K merge_friend_extras) (apply2 #3 data)); ); fun register_codatatype_extra fpT_name extra = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra)))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra)))); fun codatatype_extra_of ctxt = Symtab.lookup (#1 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_codatatype_extra (Proof_Context.theory_of ctxt)); fun all_codatatype_extras_of ctxt = Symtab.dest (#1 (Data.get (Context.Proof ctxt))); fun register_coinduct_extra fpT_name extra = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra)))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra)))); fun coinduct_extra_of ctxt = Symtab.lookup (#2 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_coinduct_extra (Proof_Context.theory_of ctxt)); fun register_friend_extra fun_name eq_algrho algrho_eq = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra) - (fn {eq_algrhos, algrho_eqs} => - {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos, - algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs})))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra) + (fn {eq_algrhos, algrho_eqs} => + {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos, + algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs})))); fun all_friend_extras_of ctxt = Symtab.dest (#3 (Data.get (Context.Proof ctxt))); fun coinduct_extras_of_generic context = corec_infos_of_generic context #> map (#corecUU #> dest_Const #> fst #> Symtab.lookup (#2 (Data.get context)) #> the #> transfer_coinduct_extra (Context.theory_of context)); fun get_coinduct_uptos fpT_name context = coinduct_extras_of_generic context fpT_name |> map #coinduct; fun get_cong_all_intros fpT_name context = coinduct_extras_of_generic context fpT_name |> maps (#cong_intro_pairs #> map snd); fun get_cong_intros fpT_name name context = coinduct_extras_of_generic context fpT_name |> map_filter (#cong_intro_pairs #> (fn ps => AList.lookup (op =) ps name)); fun ctr_names_of_fp_name lthy fpT_name = fpT_name |> fp_sugar_of lthy |> the |> #fp_ctr_sugar |> #ctr_sugar |> #ctrs |> map (Long_Name.base_name o name_of_ctr); fun register_coinduct_dynamic_base fpT_name lthy = let val fp_b = Binding.name (Long_Name.base_name fpT_name) in lthy |> fold Local_Theory.add_thms_dynamic ((mk_fp_binding fp_b coinduct_uptoN, get_coinduct_uptos fpT_name) :: map (fn N => let val N = cong_N ^ N in (mk_fp_binding fp_b N, get_cong_intros fpT_name N) end) ([baseN, reflN, symN, transN] @ ctr_names_of_fp_name lthy fpT_name)) |> Local_Theory.add_thms_dynamic (mk_fp_binding fp_b cong_introsN, get_cong_all_intros fpT_name) end; fun register_coinduct_dynamic_friend fpT_name friend_name = let val fp_b = Binding.name (Long_Name.base_name fpT_name); val friend_base_name = cong_N ^ Long_Name.base_name friend_name; in Local_Theory.add_thms_dynamic (mk_fp_binding fp_b friend_base_name, get_cong_intros fpT_name friend_base_name) end; fun derive_case_dtor ctxt fpT_name = let val thy = Proof_Context.theory_of ctxt; val SOME ({fp_res_index, fp_res = {dtors = dtors0, dtor_ctors, ...}, absT_info = {rep = rep0, abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {casex, exhaust, case_thms, ...}, ...}, ...}) = fp_sugar_of ctxt fpT_name; val (f_Ts, Type (_, [fpT as Type (_, As), _])) = strip_fun_type (fastype_of casex); val x_Tss = map binder_types f_Ts; val (((u, fs), xss), _) = ctxt |> yield_singleton (mk_Frees "y") fpT ||>> mk_Frees "f" f_Ts ||>> mk_Freess "x" x_Tss; val dtor0 = nth dtors0 fp_res_index; val dtor = mk_dtor As dtor0; val u' = dtor $ u; val absT = fastype_of u'; val rep = mk_rep absT rep0; val goal = mk_Trueprop_eq (list_comb (casex, fs) $ u, mk_case_absumprod absT rep fs xss xss $ u') |> Raw_Simplifier.rewrite_term thy @{thms comp_def[THEN eq_reflection]} []; val vars = map (fst o dest_Free) (u :: fs); val dtor_ctor = nth dtor_ctors fp_res_index; in Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_case_dtor_tac ctxt u abs_inverse dtor_ctor ctr_defs exhaust case_thms) |> Thm.close_derivation \<^here> end; fun derive_case_trivial ctxt fpT_name = let val SOME {casex, exhaust, case_thms, ...} = ctr_sugar_of ctxt fpT_name; val Type (_, As0) = domain_type (body_fun_type (fastype_of casex)); val (As, _) = ctxt |> mk_TFrees' (map Type.sort_of_atyp As0); val fpT = Type (fpT_name, As); val (var_name, ()) = singleton (Variable.variant_frees ctxt []) ("x", ()); val var = Free (var_name, fpT); val goal = mk_Trueprop_eq (expand_to_ctr_term ctxt fpT var, var); val exhaust' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt var)] exhaust; in Goal.prove_sorry ctxt [var_name] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt exhaust') THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt case_thms THEN ALLGOALS (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; fun mk_abs_rep_transfers ctxt fpT_name = [mk_abs_transfer ctxt fpT_name, mk_rep_transfer ctxt fpT_name] handle Fail _ => []; fun ensure_codatatype_extra fpT_name ctxt = (case codatatype_extra_of ctxt fpT_name of NONE => let val abs_rep_transfers = mk_abs_rep_transfers ctxt fpT_name in ctxt |> register_codatatype_extra fpT_name {case_dtor = derive_case_dtor ctxt fpT_name, case_trivial = derive_case_trivial ctxt fpT_name, abs_rep_transfers = abs_rep_transfers} |> set_transfer_rule_attrs abs_rep_transfers end | SOME {abs_rep_transfers, ...} => ctxt |> set_transfer_rule_attrs abs_rep_transfers); fun setup_base fpT_name = register_coinduct_dynamic_base fpT_name #> ensure_codatatype_extra fpT_name; fun is_set ctxt (const_name, T) = (case T of Type (\<^type_name>\fun\, [Type (fpT_name, _), Type (\<^type_name>\set\, [_])]) => (case bnf_of ctxt fpT_name of SOME bnf => exists (fn Const (s, _) => s = const_name | _ => false) (sets_of_bnf bnf) | NONE => false) | _ => false); fun case_eq_if_thms_of_term ctxt t = let val ctr_sugars = map_filter (ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in maps #case_eq_ifs ctr_sugars end; fun all_algrho_eqs_of ctxt = maps (#algrho_eqs o snd) (all_friend_extras_of ctxt); fun derive_code ctxt inner_fp_simps goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_thm, ...} fun_t fun_def = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_code_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs corecUU_thm all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps inner_fp_simps fun_def)) |> Thm.close_derivation \<^here> end; fun derive_unique ctxt phi code_goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} fpT_name eq_corecUU = let val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for lhs rhs\\ = code_goal; val (fun_t, args) = strip_comb lhs; val closed_rhs = fold_rev lambda args rhs; val fun_T = fastype_of fun_t; val num_args = num_binder_types fun_T; val f = Free (singleton (Variable.variant_frees ctxt []) ("f", fun_T)); val is_self_call = curry (op aconv) fun_t; val has_self_call = exists_subterm is_self_call; fun fify args (t $ u) = fify (u :: args) t $ fify [] u | fify _ (Abs (s, T, t)) = Abs (s, T, fify [] t) | fify args t = if t = fun_t andalso not (exists has_self_call args) then f else t; val goal = Logic.mk_implies (mk_Trueprop_eq (f, fify [] closed_rhs), mk_Trueprop_eq (f, fun_t)) |> Morphism.term phi; in Goal.prove_sorry ctxt [fst (dest_Free f)] [] goal (fn {context = ctxt, prems = _} => mk_unique_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique eq_corecUU) |> Thm.close_derivation \<^here> end; fun derive_last_disc ctxt fcT_name = let val SOME {T = fcT, discs, exhaust, disc_thmss, ...} = ctr_sugar_of ctxt fcT_name; val (u, _) = ctxt |> yield_singleton (mk_Frees "x") fcT; val udiscs = map (rapp u) discs; val (not_udiscs, last_udisc) = split_last udiscs |>> map HOLogic.mk_not; val goal = mk_Trueprop_eq (last_udisc, foldr1 HOLogic.mk_conj not_udiscs); in Goal.prove_sorry ctxt [fst (dest_Free u)] [] goal (fn {context = ctxt, prems = _} => mk_last_disc_tac ctxt u exhaust (flat disc_thmss)) |> Thm.close_derivation \<^here> end; fun derive_eq_algrho ctxt {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} ({algrho = algrho0, dtor_algrho, ...} : friend_info) fun_t k_T code_goal const_transfers rho_def eq_corecUU = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, Ts)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {fp_res_index, fp_res, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_dtor, ...} = codatatype_extra_of ctxt fpT_name; val fp_nesting_Ts = map T_of_bnf fp_nesting_bnfs; fun is_nullary_disc_def (\<^Const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _))) = true | is_nullary_disc_def (Const (\<^const_name>\Pure.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) = true | is_nullary_disc_def _ = false; val dtor_ctor = nth (#dtor_ctors fp_res) fp_res_index; val ctor_iff_dtor = #ctor_iff_dtor fp_ctr_sugar; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val nullary_disc_defs = filter (is_nullary_disc_def o Thm.prop_of) (#disc_defs ctr_sugar); val disc_sel_eq_cases = #disc_eq_cases ctr_sugar @ #sel_defs ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; fun add_tnameT (Type (s, Ts)) = insert (op =) s #> fold add_tnameT Ts | add_tnameT _ = I; fun map_disc_sels'_of s = (case fp_sugar_of ctxt s of SOME {fp_bnf_sugar = {map_disc_iffs, map_selss, ...}, ...} => let val map_selss' = if length map_selss <= 1 then map_selss else map (map (unfold_thms ctxt (no_refl [derive_last_disc ctxt s]))) map_selss; in map_disc_iffs @ flat map_selss' end | NONE => []); fun mk_const_pointful_natural const_transfer = SOME (mk_pointful_natural_from_transfer ctxt const_transfer) handle UNNATURAL () => NONE; val const_pointful_natural_opts = map mk_const_pointful_natural const_transfers; val const_pointful_naturals = map_filter I const_pointful_natural_opts; val fp_nesting_k_T_names = fold add_tnameT (k_T :: fp_nesting_Ts) []; val fp_nesting_k_map_disc_sels' = maps map_disc_sels'_of fp_nesting_k_T_names; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val ctor = nth (#ctors fp_res) fp_res_index; val abs = #abs absT_info; val rep = #rep absT_info; val algrho = mk_ctr Ts algrho0; val goal = mk_Trueprop_eq (fun_t, abs_curried_balanced arg_Ts algrho); fun const_of_transfer thm = (case Thm.prop_of thm of \<^Const>\Trueprop\ $ (_ $ cst $ _) => cst); val eq_algrho = Goal.prove (*no sorry*) ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_algrho_tac ctxt fpsig_nesting_maps abs rep ctor ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident dtor_ctor ctor_iff_dtor ctr_defs nullary_disc_defs disc_sel_eq_cases case_dtor case_eq_ifs const_pointful_naturals fp_nesting_k_map_disc_sels' rho_def dtor_algrho corecUU_unique eq_corecUU all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps) |> Thm.close_derivation \<^here> handle e as ERROR _ => (case filter (is_none o snd) (const_transfers ~~ const_pointful_natural_opts) of [] => Exn.reraise e | thm_nones => error ("Failed to state naturality property for " ^ commas (map (Syntax.string_of_term ctxt o const_of_transfer o fst) thm_nones))); val algrho_eq = eq_algrho RS (mk_curry_uncurryN_balanced ctxt num_args RS iffD2) RS sym; in (eq_algrho, algrho_eq) end; fun prime_rho_transfer_goal ctxt fpT_name rho_def goal = let val thy = Proof_Context.theory_of ctxt; val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; val simps = rel_def_of_bnf pre_bnf :: rho_transfer_simps; val fold_rho = unfold_thms ctxt [rho_def RS @{thm symmetric}]; fun derive_unprimed rho_transfer' = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt simps THEN HEADGOAL (rtac ctxt rho_transfer'))) |> Thm.close_derivation \<^here>; val goal' = Raw_Simplifier.rewrite_term thy simps [] goal; in if null abs_rep_transfers then (goal', derive_unprimed #> fold_rho) else (goal, fold_rho) end; fun derive_rho_transfer_folded ctxt fpT_name const_transfers rho_def goal = let val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_rho_transfer_tac ctxt (null abs_rep_transfers) (rel_def_of_bnf pre_bnf) const_transfers)) |> unfold_thms ctxt [rho_def RS @{thm symmetric}] |> Thm.close_derivation \<^here> end; fun mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong alg = let val xy_Ts = binder_types (fastype_of alg); val ((xs, ys), _) = ctxt |> mk_Frees "x" xy_Ts ||>> mk_Frees "y" xy_Ts; fun mk_prem xy_T x y = build_rel [] ctxt [fpT] [] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T) (xy_T, xy_T) $ x $ y; val prems = @{map 3} mk_prem xy_Ts xs ys; val concl = Rcong $ list_comb (alg, xs) $ list_comb (alg, ys); in Logic.list_implies (map HOLogic.mk_Trueprop prems, HOLogic.mk_Trueprop concl) end; fun derive_cong_ctr_intros ctxt cong_ctor_intro = let val \<^Const_>\Pure.imp\ $ _ $ (\<^Const_>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ (ctor $ _))) = Thm.prop_of cong_ctor_intro; val fpT as Type (fpT_name, fp_argTs) = range_type (fastype_of ctor); val SOME {pre_bnf, absT_info = {abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {ctrs = ctrs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val ctrs = map (mk_ctr fp_argTs) ctrs0; val pre_rel_def = rel_def_of_bnf pre_bnf; fun prove ctr_def goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt ctr_def [pre_rel_def, abs_inverse] cong_ctor_intro)) |> Thm.close_derivation \<^here>; val goals = map (mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong) ctrs; in map2 prove ctr_defs goals end; fun derive_cong_friend_intro ctxt cong_algrho_intro = let val \<^Const_>\Pure.imp\ $ _ $ (\<^Const_>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ ((algrho as Const (algrho_name, _)) $ _))) = Thm.prop_of cong_algrho_intro; val fpT as Type (_, fp_argTs) = range_type (fastype_of algrho); fun has_algrho (\<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ rhs)) = fst (dest_Const (head_of (strip_abs_body rhs))) = algrho_name; val eq_algrho :: _ = maps (filter (has_algrho o Thm.prop_of) o #eq_algrhos o snd) (all_friend_extras_of ctxt); val \<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ friend0 $ _) = Thm.prop_of eq_algrho; val friend = mk_ctr fp_argTs friend0; val goal = mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong friend; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt eq_algrho [] cong_algrho_intro)) |> Thm.close_derivation \<^here> end; fun derive_cong_intros lthy ctr_names friend_names ({cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros, ...} : dtor_coinduct_info) = let val cong_ctor_intro :: cong_algrho_intros = rev cong_alg_intros; val names = map (prefix cong_N) ([baseN, reflN, symN, transN] @ ctr_names @ friend_names); val thms = [cong_base, cong_refl, cong_sym, cong_trans] @ derive_cong_ctr_intros lthy cong_ctor_intro @ map (derive_cong_friend_intro lthy) cong_algrho_intros; in names ~~ thms end; fun derive_coinduct ctxt (fpT as Type (fpT_name, fpT_args)) dtor_coinduct = let val thy = Proof_Context.theory_of ctxt; val \<^Const_>\Pure.imp\ $ (\<^Const_>\Trueprop\ $ (_ $ Abs (_, _, _ $ Abs (_, _, \<^Const_>\implies\ $ _ $ (_ $ (cong0 $ _) $ _ $ _))))) $ _ = Thm.prop_of dtor_coinduct; val SOME {X as TVar ((X_s, _), _), fp_res = {dtor_ctors, ...}, pre_bnf, absT_info = {abs_inverse, ...}, live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss, ctr_defs, ctr_sugar = ctr_sugar0 as {T = Type (_, T0_args), ctrs = ctrs0, discs = discs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val n = length ctrXs_Tss; val ms = map length ctrXs_Tss; val X' = TVar ((X_s, maxidx_of_typ fpT + 1), \<^sort>\type\); val As_rho = tvar_subst thy T0_args fpT_args; val substXAT = Term.typ_subst_TVars As_rho o Tsubst X X'; val substXA = Term.subst_TVars As_rho o substT X X'; val phi = Morphism.typ_morphism "BNF" substXAT $> Morphism.term_morphism "BNF" substXA; fun mk_applied_cong arg = enforce_type ctxt domain_type (fastype_of arg) cong0 $ arg; val thm = derive_coinduct_thms_for_types ctxt false mk_applied_cong [pre_bnf] dtor_coinduct dtor_ctors live_nesting_bnfs [fpT] [substXAT X] [map (map substXAT) ctrXs_Tss] [n] [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0] |> map snd |> the_single; val (attrs, _) = mk_coinduct_attrs [fpT] [ctrs0] [discs0] [ms]; in (thm, attrs) end; type explore_parameters = {bound_Us: typ list, bound_Ts: typ list, U: typ, T: typ}; fun update_UT {bound_Us, bound_Ts, ...} U T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = U, T = T}; fun explore_nested lthy explore {bound_Us, bound_Ts, U, T} t = let fun build_simple (T, U) = if T = U then \<^term>\%y. y\ else Bound 0 |> explore {bound_Us = T :: bound_Us, bound_Ts = T :: bound_Ts, U = U, T = T} |> (fn t => Abs (Name.uu, T, t)); in betapply (build_map lthy [] [] build_simple (T, U), t) end; fun add_boundvar t = betapply (incr_boundvars 1 t, Bound 0); fun explore_fun (arg_U :: arg_Us) explore {bound_Us, bound_Ts, U, T} t = let val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t) in add_boundvar t |> explore_fun arg_Us explore {bound_Us = arg_U :: bound_Us, bound_Ts = domain_type T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_U, t)) end | explore_fun [] explore params t = explore params t; fun massage_fun explore (params as {T, U, ...}) = if can dest_funT T then explore_fun [domain_type U] explore params else explore params; fun massage_star massages explore = let fun after_massage massages' t params t' = if t aconv t' then massage_any massages' params t else massage_any massages params t' and massage_any [] params t = explore params t | massage_any (massage :: massages') params t = massage (after_massage massages' t) params t; in massage_any massages end; fun massage_let explore params t = (case strip_comb t of (Const (\<^const_name>\Let\, _), [_, _]) => unfold_lets_splits t | _ => t) |> explore params; fun check_corec_equation ctxt fun_frees (lhs, rhs) = let val (fun_t, arg_ts) = strip_comb lhs; fun check_fun_name () = null fun_frees orelse member (op aconv) fun_frees fun_t orelse ill_formed_equation_head ctxt [] fun_t; fun check_no_other_frees () = (case Term.add_frees rhs [] |> map Free |> subtract (op =) (fun_frees @ arg_ts) |> find_first (not o Variable.is_fixed ctxt o fst o dest_Free) of NONE => () | SOME t => extra_variable_in_rhs ctxt [] t); in check_duplicate_variables_in_lhs ctxt [] arg_ts; check_fun_name (); check_all_fun_arg_frees ctxt [] (filter_out is_Var arg_ts); check_no_other_frees () end; fun parse_corec_equation ctxt fun_frees eq = let val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (drop_all eq)) handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eq]; val _ = check_corec_equation ctxt fun_frees (lhs, rhs); val (fun_t, arg_ts) = strip_comb lhs; val (arg_Ts, _) = strip_type (fastype_of fun_t); val added_Ts = drop (length arg_ts) arg_Ts; val free_names = mk_names (length added_Ts) "x" ~~ added_Ts; val free_args = Variable.variant_frees ctxt [rhs, lhs] free_names |> map Free; in (arg_ts @ free_args, list_comb (rhs, free_args)) end; fun morph_views phi (code, ctrs, discs, disc_iffs, sels) = (Morphism.term phi code, map (Morphism.term phi) ctrs, map (Morphism.term phi) discs, map (Morphism.term phi) disc_iffs, map (Morphism.term phi) sels); fun generate_views ctxt eq fun_t (lhs_free_args, rhs) = let val lhs = list_comb (fun_t, lhs_free_args); val T as Type (T_name, Ts) = fastype_of rhs; val SOME {fp_ctr_sugar = {ctr_sugar = {ctrs = ctrs0, discs = discs0, selss = selss0, ...}, ...}, ...} = fp_sugar_of ctxt T_name; val ctrs = map (mk_ctr Ts) ctrs0; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val code_view = drop_all eq; fun can_case_expand t = not (can (dest_ctr ctxt T_name) t); fun generate_raw_views conds t raw_views = let fun analyse (ctr :: ctrs) (disc :: discs) ctr' = if ctr = ctr' then (conds, disc, ctr) else analyse ctrs discs ctr'; in (analyse ctrs discs (fst (strip_comb t))) :: raw_views end; fun generate_disc_views raw_views = if length discs = 1 then ([], []) else let fun collect_condss_disc condss [] _ = condss | collect_condss_disc condss ((conds, disc', _) :: raw_views) disc = collect_condss_disc (condss |> disc = disc' ? cons conds) raw_views disc; val grouped_disc_views = discs |> map (collect_condss_disc [] raw_views) |> curry (op ~~) (map (fn disc => disc $ lhs) discs); fun mk_disc_iff_props props [] = props | mk_disc_iff_props _ ((lhs, \<^Const_>\True\) :: _) = [lhs] | mk_disc_iff_props props ((lhs, rhs) :: views) = mk_disc_iff_props ((HOLogic.mk_eq (lhs, rhs)) :: props) views; in (grouped_disc_views |> map swap, grouped_disc_views |> map (apsnd (s_dnf #> mk_conjs)) |> mk_disc_iff_props [] |> map (fn eq => ([[]], eq))) end; fun generate_ctr_views raw_views = let fun collect_condss_ctr condss [] _ = condss | collect_condss_ctr condss ((conds, _, ctr') :: raw_views) ctr = collect_condss_ctr (condss |> ctr = ctr' ? cons conds) raw_views ctr; fun mk_ctr_eq ctr_sels ctr = let fun extract_arg n sel _(*bound_Ts*) fun_t arg_ts = if ctr = fun_t then nth arg_ts n else let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern (range_type (fastype_of sel)) end; in ctr_sels |> map_index (uncurry extract_arg) |> map (fn extract => massage_corec_code_rhs ctxt extract [] rhs) |> curry list_comb ctr |> curry HOLogic.mk_eq lhs end; fun remove_condss_if_alone [(_, concl)] = [([[]], concl)] | remove_condss_if_alone views = views; in ctrs |> `(map (collect_condss_ctr [] raw_views)) ||> map2 mk_ctr_eq selss |> op ~~ |> filter_out (null o fst) |> remove_condss_if_alone end; fun generate_sel_views raw_views only_one_ctr = let fun mk_sel_positions sel = let fun get_sel_position _ [] = NONE | get_sel_position i (sel' :: sels) = if sel = sel' then SOME i else get_sel_position (i + 1) sels; in ctrs ~~ map (get_sel_position 0) selss |> map_filter (fn (ctr, pos_opt) => if is_some pos_opt then SOME (ctr, the pos_opt) else NONE) end; fun collect_sel_condss0 condss [] _ = condss | collect_sel_condss0 condss ((conds, _, ctr) :: raw_views) sel_positions = let val condss' = condss |> is_some (AList.lookup (op =) sel_positions ctr) ? cons conds in collect_sel_condss0 condss' raw_views sel_positions end; val collect_sel_condss = if only_one_ctr then K [[]] else collect_sel_condss0 [] raw_views; fun mk_sel_rhs sel_positions sel = let val sel_T = range_type (fastype_of sel); fun extract_sel_value _(*bound_Ts*) fun_t arg_ts = (case AList.lookup (op =) sel_positions fun_t of SOME n => nth arg_ts n | NONE => let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern sel_T end); in massage_corec_code_rhs ctxt extract_sel_value [] rhs end; val ordered_sels = distinct (op =) (flat selss); val sel_positionss = map mk_sel_positions ordered_sels; val sel_rhss = map2 mk_sel_rhs sel_positionss ordered_sels; val sel_lhss = map (rapp lhs o mk_disc_or_sel Ts) ordered_sels; val sel_condss = map collect_sel_condss sel_positionss; fun is_undefined (Const (\<^const_name>\undefined\, _)) = true | is_undefined _ = false; in sel_condss ~~ (sel_lhss ~~ sel_rhss) |> filter_out (is_undefined o snd o snd) |> map (apsnd HOLogic.mk_eq) end; fun mk_atomic_prop fun_args (condss, concl) = (Logic.list_all (map dest_Free fun_args, abstract_over_list fun_args (Logic.list_implies (map HOLogic.mk_Trueprop (s_dnf condss), HOLogic.mk_Trueprop concl)))); val raw_views = rhs |> massage_let_if_case ctxt (K false) (fn _(*bound_Ts*) => fn t => t |> can_case_expand t ? expand_to_ctr_term ctxt T) (K (K ())) (K I) [] |> (fn expanded_rhs => fold_rev_let_if_case ctxt generate_raw_views [] expanded_rhs []) |> rev; val (disc_views, disc_iff_views) = generate_disc_views raw_views; val ctr_views = generate_ctr_views raw_views; val sel_views = generate_sel_views raw_views (length ctr_views = 1); val mk_props = filter_out (null o fst) #> map (mk_atomic_prop lhs_free_args); in (code_view, mk_props ctr_views, mk_props disc_views, mk_props disc_iff_views, mk_props sel_views) end; fun find_all_associated_types [] _ = [] | find_all_associated_types ((Type (_, Ts1), Type (_, Ts2)) :: TTs) T = find_all_associated_types ((Ts1 ~~ Ts2) @ TTs) T | find_all_associated_types ((T1, T2) :: TTs) T = find_all_associated_types TTs T |> T1 = T ? cons T2; fun as_member_of tab = try dest_Const #> Option.mapPartial (fst #> Symtab.lookup tab); fun extract_rho_from_equation ({ctr_guards, inner_buffer = {Oper, VLeaf, CLeaf, ctr_wrapper, friends}, ...}, {pattern_ctrs, discs, sels, it, mk_case}) b version Y preT ssig_T friend_tm (lhs_args, rhs) lthy = let val thy = Proof_Context.theory_of lthy; val res_T = fastype_of rhs; val YpreT = HOLogic.mk_prodT (Y, preT); fun fpT_to new_T T = if T = res_T then new_T else (case T of Type (s, Ts) => Type (s, map (fpT_to new_T) Ts) | _ => T); fun build_params bound_Us bound_Ts T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = T, T = T}; fun typ_before explore {bound_Us, bound_Ts, ...} t = explore (build_params bound_Us bound_Ts (fastype_of1 (bound_Ts, t))) t; val is_self_call = curry (op aconv) friend_tm; val has_self_call = Term.exists_subterm is_self_call; fun has_res_T bound_Ts t = fastype_of1 (bound_Ts, t) = res_T; fun contains_res_T (Type (s, Ts)) = s = fst (dest_Type res_T) orelse exists contains_res_T Ts | contains_res_T _ = false; val res_T_lhs_args = filter (exists_type contains_res_T) lhs_args; val is_res_T_lhs_arg = member (op =) res_T_lhs_args; fun is_constant t = not (Term.exists_subterm is_res_T_lhs_arg t orelse has_self_call t orelse loose_bvar (t, 0)); fun is_nested_type T = T <> res_T andalso T <> YpreT andalso T <> ssig_T; val is_valid_case_argumentT = not o member (op =) [Y, ssig_T]; fun is_same_type_constr (Type (s, _)) (Type (s', _)) = (s = s') | is_same_type_constr _ _ = false; exception NO_ENCAPSULATION of unit; val parametric_consts = Unsynchronized.ref []; (* We are assuming that set functions are marked with "[transfer_rule]" (cf. the "transfer" plugin). Otherwise, the "eq_algrho" tactic might fail. *) fun is_special_parametric_const (x as (s, _)) = s = \<^const_name>\id\ orelse is_set lthy x; fun add_parametric_const s general_T T U = let fun tupleT_of_funT T = let val (Ts, T) = strip_type T in mk_tupleT_balanced (Ts @ [T]) end; fun funT_of_tupleT n = dest_tupleT_balanced (n + 1) #> split_last #> op --->; val m = num_binder_types general_T; val param1_T = Type_Infer.paramify_vars general_T; val param2_T = Type_Infer.paramify_vars param1_T; val deadfixed_T = build_map lthy [] [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T)) |> singleton (Type_Infer_Context.infer_types lthy) |> singleton (Type_Infer.fixate lthy false) |> type_of |> dest_funT |-> generalize_types 1 |> funT_of_tupleT m; val j = maxidx_of_typ deadfixed_T + 1; fun varifyT (Type (s, Ts)) = Type (s, map varifyT Ts) | varifyT (TFree (s, T)) = TVar ((s, j), T) | varifyT T = T; val dedvarified_T = varifyT deadfixed_T; val new_vars = Sign.typ_match thy (dedvarified_T, T) Vartab.empty |> Vartab.dest |> filter (curry (op =) j o snd o fst) |> Vartab.make; val deadinstantiated_T = map_atyps (Type.devar new_vars) dedvarified_T; val final_T = if Sign.typ_instance thy (U, deadinstantiated_T) then deadfixed_T else general_T; in parametric_consts := insert (op =) (s, final_T) (!parametric_consts) end; fun encapsulate (params as {U, T, ...}) t = if U = T then t else if T = Y then VLeaf $ t else if T = res_T then CLeaf $ t else if T = YpreT then it $ t else if is_nested_type T andalso is_same_type_constr T U then explore_nested lthy encapsulate params t else raise NO_ENCAPSULATION (); fun build_function_after_encapsulation fun_t fun_t' (params as {bound_Us, ...}) arg_ts arg_ts' = let val arg_Us' = fst (strip_typeN (length arg_ts) (fastype_of1 (bound_Us, fun_t'))); fun the_or_error arg NONE = error ("Illegal argument " ^ quote (Syntax.string_of_term lthy arg) ^ " to " ^ quote (Syntax.string_of_term lthy fun_t)) | the_or_error _ (SOME arg') = arg'; in arg_ts' |> `(map (curry fastype_of1 bound_Us)) |>> map2 (update_UT params) arg_Us' |> op ~~ |> map (try (uncurry encapsulate)) |> map2 the_or_error arg_ts |> curry list_comb fun_t' end; fun rebuild_function_after_exploration old_fn new_fn explore params arg_ts = arg_ts |> map (typ_before explore params) |> build_function_after_encapsulation old_fn new_fn params arg_ts; fun update_case Us U casex = let val Type (T_name, _) = domain_type (snd (strip_fun_type (fastype_of casex))); val SOME {fp_ctr_sugar = {ctr_sugar = {T = Type (_, Ts), casex, ...}, ...}, ...} = fp_sugar_of lthy T_name; val T = body_type (fastype_of casex); in Term.subst_atomic_types ((T :: Ts) ~~ (U :: Us)) casex end; fun deduce_according_type default_T [] = default_T | deduce_according_type default_T Ts = (case distinct (op =) Ts of U :: [] => U | _ => fpT_to ssig_T default_T); fun massage_if explore_cond explore (params as {bound_Us, bound_Ts, ...}) t = (case strip_comb t of (const as Const (\<^const_name>\If\, _), obj :: (branches as [_, _])) => (case List.partition Term.is_dummy_pattern (map (explore params) branches) of (dummy_branch' :: _, []) => dummy_branch' | (_, [branch']) => branch' | (_, branches') => let val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (fastype_of1 (bound_Ts, hd branches)) brancheUs; val const_obj' = (If_const U, obj) ||> explore_cond (update_UT params \<^typ>\bool\ \<^typ>\bool\) |> op $; in build_function_after_encapsulation (const $ obj) const_obj' params branches branches' end) | _ => explore params t); fun massage_map explore (params as {bound_Us, bound_Ts, T = Type (T_name, Ts), ...}) (t as func $ mapped_arg) = if is_self_call (head_of func) then explore params t else (case try (dest_map lthy T_name) func of SOME (map_tm, fs) => let val n = length fs; val mapped_arg' = mapped_arg |> `(curry fastype_of1 bound_Ts) |>> build_params bound_Us bound_Ts |-> explore; in (case fastype_of1 (bound_Us, mapped_arg') of Type (U_name, Us0) => if U_name = T_name then let val Us = map (fpT_to ssig_T) Us0; val temporary_map = map_tm |> mk_map n Us Ts; val map_fn_Ts = fastype_of #> strip_fun_type #> fst; val binder_Uss = map_fn_Ts temporary_map |> map (map (fpT_to ssig_T) o binder_types); val fun_paramss = map_fn_Ts (head_of func) |> map (build_params bound_Us bound_Ts); val fs' = fs |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss; val SOME bnf = bnf_of lthy T_name; val Type (_, bnf_Ts) = T_of_bnf bnf; val typ_alist = lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs'; val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts); val map_tm' = map_tm |> mk_map n Us Us'; in build_function_after_encapsulation func (list_comb (map_tm', fs')) params [mapped_arg] [mapped_arg'] end else explore params t | _ => explore params t) end | NONE => explore params t) | massage_map explore params t = explore params t; fun massage_comp explore (params as {bound_Us, ...}) t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => let val args' = map (typ_before explore params) args; val f2' = typ_before (explore_fun (map (curry fastype_of1 bound_Us) args') explore) params f2; val f1' = typ_before (explore_fun [range_type (fastype_of1 (bound_Us, f2'))] explore) params f1; in betapply (f1', list_comb (f2', args')) end | _ => explore params t); fun massage_ctr explore (params as {T = T as Type (s, Ts), bound_Us, ...}) t = if T <> res_T then (case try (dest_ctr lthy s) t of SOME (ctr, args) => let val args' = map (typ_before explore params) args; val SOME {T = Type (_, ctr_Ts), ...} = ctr_sugar_of lthy s; val temp_ctr = mk_ctr ctr_Ts ctr; val argUs = map (curry fastype_of1 bound_Us) args'; val typ_alist = binder_types (fastype_of temp_ctr) ~~ argUs; val Us = ctr_Ts |> map (find_all_associated_types typ_alist) |> map2 deduce_according_type Ts; val ctr' = mk_ctr Us ctr; in build_function_after_encapsulation ctr ctr' params args args' end | NONE => explore params t) else explore params t | massage_ctr explore params t = explore params t; fun const_of [] _ = NONE | const_of ((sel as Const (s1, _)) :: r) (const as Const (s2, _)) = if s1 = s2 then SOME sel else const_of r const | const_of _ _ = NONE; fun massage_disc explore (params as {T, bound_Us, bound_Ts, ...}) t = (case (strip_comb t, T = \<^typ>\bool\) of ((fun_t, arg :: []), true) => let val arg_T = fastype_of1 (bound_Ts, arg) in if arg_T <> res_T then (case arg_T |> try (fst o dest_Type) |> Option.mapPartial (ctr_sugar_of lthy) of SOME {discs, T = Type (_, Ts), ...} => (case const_of discs fun_t of SOME disc => let val arg' = arg |> typ_before explore params; val Type (_, Us) = fastype_of1 (bound_Us, arg'); val disc' = disc |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in disc' $ arg' end | NONE => explore params t) | NONE => explore params t) else explore params t end | _ => explore params t); fun massage_sel explore (params as {bound_Us, bound_Ts, ...}) t = let val (fun_t, args) = strip_comb t in if args = [] then explore params t else let val T = fastype_of1 (bound_Ts, hd args) in (case (Option.mapPartial (ctr_sugar_of lthy) (try (fst o dest_Type) T), T <> res_T) of (SOME {selss, T = Type (_, Ts), ...}, true) => (case const_of (flat selss) fun_t of SOME sel => let val args' = args |> map (typ_before explore params); val Type (_, Us) = fastype_of1 (bound_Us, hd args'); val sel' = sel |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in build_function_after_encapsulation sel sel' params args args' end | NONE => explore params t) | _ => explore params t) end end; fun massage_equality explore (params as {bound_Us, bound_Ts, ...}) (t as Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = let val check_is_VLeaf = not o (Term.exists_subterm (fn t => t aconv CLeaf orelse t aconv Oper)); fun try_pattern_matching (fun_t, arg_ts) t = (case as_member_of pattern_ctrs fun_t of SOME (disc, sels) => let val t' = typ_before explore params t in if fastype_of1 (bound_Us, t') = YpreT then let val arg_ts' = map (typ_before explore params) arg_ts; val sels_t' = map (fn sel => betapply (sel, t')) sels; val Ts = map (curry fastype_of1 bound_Us) arg_ts'; val Us = map (curry fastype_of1 bound_Us) sels_t'; val arg_ts' = map2 encapsulate (map2 (update_UT params) Us Ts) arg_ts'; in if forall check_is_VLeaf arg_ts' then SOME (Library.foldl1 HOLogic.mk_conj (betapply (disc, t') :: (map HOLogic.mk_eq (arg_ts' ~~ sels_t')))) else NONE end else NONE end | NONE => NONE); in (case try_pattern_matching (strip_comb t1) t2 of SOME cond => cond | NONE => (case try_pattern_matching (strip_comb t2) t1 of SOME cond => cond | NONE => let val T = fastype_of1 (bound_Ts, t1); val params' = build_params bound_Us bound_Ts T; val t1' = explore params' t1; val t2' = explore params' t2; in if fastype_of1 (bound_Us, t1') = T andalso fastype_of1 (bound_Us, t2') = T then HOLogic.mk_eq (t1', t2') else error ("Unsupported condition: " ^ quote (Syntax.string_of_term lthy t)) end)) end | massage_equality explore params t = explore params t; fun infer_types (TVar _) (TVar _) = [] | infer_types (U as TVar _) T = [(U, T)] | infer_types (Type (s', Us)) (Type (s, Ts)) = if s' = s then flat (map2 infer_types Us Ts) else [] | infer_types _ _ = []; fun group_by_fst associations [] = associations | group_by_fst associations ((a, b) :: r) = group_by_fst (add_association a b associations) r and add_association a b [] = [(a, [b])] | add_association a b ((c, d) :: r) = if a = c then (c, b :: d) :: r else (c, d) :: (add_association a b r); fun new_TVar known_TVars = Name.invent_list (map (fst o fst o dest_TVar) known_TVars) "x" 1 |> (fn [s] => TVar ((s, 0), [])); fun instantiate_type inferred_types = Term.typ_subst_TVars (map (apfst (fst o dest_TVar)) inferred_types); fun chose_unknown_TVar (T as TVar _) = SOME T | chose_unknown_TVar (Type (_, Ts)) = fold (curry merge_options) (map chose_unknown_TVar Ts) NONE | chose_unknown_TVar _ = NONE; (* The function under definition might not be defined yet when this is queried. *) fun maybe_const_type ctxt (s, T) = Sign.const_type (Proof_Context.theory_of ctxt) s |> the_default T; fun massage_const polymorphic explore (params as {bound_Us, ...}) t = let val (fun_t, arg_ts) = strip_comb t in (case fun_t of Const (fun_x as (s, fun_T)) => let val general_T = if polymorphic then maybe_const_type lthy fun_x else fun_T in if fun_t aconv friend_tm orelse contains_res_T (body_type general_T) orelse is_constant t then explore params t else let val inferred_types = infer_types general_T fun_T; fun prepare_skeleton [] _ = [] | prepare_skeleton ((T, U) :: inferred_types) As = let fun schematize_res_T U As = if U = res_T then let val A = new_TVar As in (A, A :: As) end else (case U of Type (s, Us) => fold_map schematize_res_T Us As |>> curry Type s | _ => (U, As)); val (U', As') = schematize_res_T U As; in (T, U') :: (prepare_skeleton inferred_types As') end; val inferred_types' = prepare_skeleton inferred_types (map fst inferred_types); val skeleton_T = instantiate_type inferred_types' general_T; fun explore_if_possible (exp_arg as (_, true)) _ = exp_arg | explore_if_possible (exp_arg as (arg, false)) arg_T = if exists (exists_subtype is_TVar) (binder_types arg_T) then exp_arg else (typ_before (explore_fun (binder_types arg_T) explore) params arg, true); fun collect_inferred_types [] _ = [] | collect_inferred_types ((arg, explored) :: exp_arg_ts) (arg_T :: arg_Ts) = (if explored then infer_types arg_T (fastype_of1 (bound_Us, arg)) else []) @ collect_inferred_types exp_arg_ts arg_Ts; fun propagate exp_arg_ts skeleton_T = let val arg_gen_Ts = binder_types skeleton_T; val exp_arg_ts = map2 explore_if_possible exp_arg_ts arg_gen_Ts; val inferred_types = collect_inferred_types exp_arg_ts arg_gen_Ts |> group_by_fst [] |> map (apsnd (deduce_according_type ssig_T)); in (exp_arg_ts, instantiate_type inferred_types skeleton_T) end; val remaining_to_be_explored = filter_out snd #> length; fun try_exploring_args exp_arg_ts skeleton_T = let val n = remaining_to_be_explored exp_arg_ts; val (exp_arg_ts', skeleton_T') = propagate exp_arg_ts skeleton_T; val n' = remaining_to_be_explored exp_arg_ts'; fun try_instantiating A T = try (try_exploring_args exp_arg_ts') (instantiate_type [(A, T)] skeleton_T'); in if n' = 0 then SOME (exp_arg_ts', skeleton_T') else if n = n' then if exists_subtype is_TVar skeleton_T' then let val SOME A = chose_unknown_TVar skeleton_T' in (case try_instantiating A ssig_T of SOME result => result | NONE => (case try_instantiating A YpreT of SOME result => result | NONE => (case try_instantiating A res_T of SOME result => result | NONE => NONE))) end else NONE else try_exploring_args exp_arg_ts' skeleton_T' end; in (case try_exploring_args (map (fn arg => (arg, false)) arg_ts) skeleton_T of SOME (exp_arg_ts, fun_U) => let val arg_ts' = map fst exp_arg_ts; val fun_t' = Const (s, fun_U); fun finish_off () = let val t' = build_function_after_encapsulation fun_t fun_t' params arg_ts arg_ts'; in if can type_of1 (bound_Us, t') then (if fun_T = fun_U orelse is_special_parametric_const (s, fun_T) then () else add_parametric_const s general_T fun_T fun_U; t') else explore params t end; in if polymorphic then finish_off () else (case try finish_off () of SOME t' => t' | NONE => explore params t) end | NONE => explore params t) end end | _ => explore params t) end; fun massage_rho explore = massage_star [massage_let, massage_if explore_cond, massage_case, massage_fun, massage_comp, massage_map, massage_ctr, massage_sel, massage_disc, massage_equality, massage_const false, massage_const true] explore and massage_case explore (params as {bound_Ts, bound_Us, ...}) t = (case strip_comb t of (casex as Const (case_x as (c, _)), args as _ :: _ :: _) => (case try strip_fun_type (maybe_const_type lthy case_x) of SOME (gen_branch_Ts, gen_body_fun_T) => let val gen_branch_ms = map num_binder_types gen_branch_Ts; val n = length gen_branch_ms; val (branches, obj_leftovers) = chop n args; in if n < length args then (case gen_body_fun_T of Type (_, [Type (T_name, _), _]) => if case_of lthy T_name = SOME (c, true) then let val brancheTs = binder_fun_types (fastype_of1 (bound_Ts, casex)); val obj_leftover_Ts = map (curry fastype_of1 bound_Ts) obj_leftovers; val obj_leftovers' = if is_constant (hd obj_leftovers) then obj_leftovers else (obj_leftover_Ts, obj_leftovers) |>> map (build_params bound_Us bound_Ts) |> op ~~ |> map (uncurry explore_inner); val obj_leftoverUs = obj_leftovers' |> map (curry fastype_of1 bound_Us); val _ = is_valid_case_argumentT (hd obj_leftoverUs) orelse error (quote (Syntax.string_of_term lthy (hd obj_leftovers)) ^ " is not a valid case argument"); val Us = obj_leftoverUs |> hd |> dest_Type |> snd; val branche_binderUss = (if hd obj_leftoverUs = YpreT then mk_case HOLogic.boolT else update_case Us HOLogic.boolT casex) |> fastype_of |> binder_fun_types |> map binder_types; val b_params = map (build_params bound_Us bound_Ts) brancheTs; val branches' = branches |> @{map 4} explore_fun branche_binderUss (replicate n explore) b_params; val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (body_type (hd brancheTs)) (map body_type brancheUs); val casex' = if hd obj_leftoverUs = YpreT then mk_case U else update_case Us U casex; in build_function_after_encapsulation casex casex' params (branches @ obj_leftovers) (branches' @ obj_leftovers') end else explore params t | _ => explore params t) else explore params t end | NONE => explore params t) | _ => explore params t) and explore_cond params t = if has_self_call t then unexpected_rec_call_in lthy [] t else explore_inner params t and explore_inner params t = massage_rho explore_inner_general params t and explore_inner_general (params as {bound_Us, bound_Ts, T, ...}) t = let val (fun_t, arg_ts) = strip_comb t in if is_constant t then t else (case (as_member_of discs fun_t, length arg_ts = 1 andalso has_res_T bound_Ts (the_single arg_ts)) of (SOME disc', true) => let val arg' = explore_inner params (the_single arg_ts); val arg_U = fastype_of1 (bound_Us, arg'); in if arg_U = res_T then fun_t $ arg' else if arg_U = YpreT then disc' $ arg' else error ("Discriminator " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | _ => (case as_member_of sels fun_t of SOME sel' => let val arg_ts' = map (explore_inner params) arg_ts; val arg_U = fastype_of1 (bound_Us, hd arg_ts'); in if arg_U = res_T then build_function_after_encapsulation fun_t fun_t params arg_ts arg_ts' else if arg_U = YpreT then build_function_after_encapsulation fun_t sel' params arg_ts arg_ts' else error ("Selector " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | NONE => (case as_member_of friends fun_t of SOME (_, friend') => rebuild_function_after_exploration fun_t friend' explore_inner params arg_ts |> curry (op $) Oper | NONE => (case as_member_of ctr_guards fun_t of SOME ctr_guard' => rebuild_function_after_exploration fun_t ctr_guard' explore_inner params arg_ts |> curry (op $) ctr_wrapper |> curry (op $) Oper | NONE => if is_Bound fun_t then rebuild_function_after_exploration fun_t fun_t explore_inner params arg_ts else if is_Free fun_t then let val fun_t' = map_types (fpT_to YpreT) fun_t in rebuild_function_after_exploration fun_t fun_t' explore_inner params arg_ts end else if T = res_T then error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this and no friend") else error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this"))))) end; fun explore_ctr params t = massage_rho explore_ctr_general params t and explore_ctr_general params t = let val (fun_t, arg_ts) = strip_comb t; val ctr_opt = as_member_of ctr_guards fun_t; in if is_some ctr_opt then rebuild_function_after_exploration fun_t (the ctr_opt) explore_inner params arg_ts else not_constructor_in_rhs lthy [] fun_t end; val rho_rhs = rhs |> explore_ctr (build_params [] [] (fastype_of rhs)) |> abs_tuple_balanced (map (map_types (fpT_to YpreT)) lhs_args) |> unfold_id_bnf_etc lthy; in lthy |> define_const false b version rhoN rho_rhs |>> pair (!parametric_consts, rho_rhs) end; fun mk_rho_parametricity_goal ctxt Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreT = Tsubst Y Z YpreT; val ssigZ_T = Tsubst Y Z ssig_T; val dead_pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssigZ_T)] dead_pre_rel; val dead_k_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreT)] dead_k_rel; val (R, _) = ctxt |> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val rho_rel = mk_rel_fun (dead_k_rel' $ mk_rel_prod R (dead_pre_rel $ R)) (dead_pre_rel' $ (dead_ssig_rel $ R)); val rho_rhsZ = substT Y Z rho_rhs; in HOLogic.mk_Trueprop (rho_rel $ rho_rhs $ rho_rhsZ) end; fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy = let fun mk_rel T bnf = let val ZT = Tsubst Y Z T; val rel_T = mk_predT [mk_pred2T Y Z, T, ZT]; in enforce_type lthy I rel_T (rel_of_bnf bnf) end; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (dead_ssig_bnf, lthy') = bnf_kill_all_but 1 ssig_bnf lthy; val dead_pre_rel = mk_rel preT dead_pre_bnf; val dead_k_rel = mk_rel k_T dead_k_bnf; val dead_ssig_rel = mk_rel ssig_T dead_ssig_bnf; val (((parametric_consts, rho_rhs), rho_data), lthy'') = extract_rho_from_equation friend_parse_info fun_b version Y preT ssig_T fun_t parsed_eq lthy'; val const_transfer_goals = map (mk_const_transfer_goal lthy'') parametric_consts; val rho_transfer_goal = mk_rho_parametricity_goal lthy'' Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs; in ((rho_data, (const_transfer_goals, rho_transfer_goal)), lthy'') end; fun explore_corec_equation ctxt could_be_friend friend fun_name fun_free {outer_buffer, ctr_guards, inner_buffer} res_T (args, rhs) = let val is_self_call = curry (op aconv) fun_free; val has_self_call = Term.exists_subterm is_self_call; val outer_ssig_T = body_type (fastype_of (#Oper outer_buffer)); fun inner_fp_of (Free (s, _)) = Free (s ^ inner_fp_suffix, mk_tupleT_balanced (map fastype_of args) --> outer_ssig_T); fun build_params bound_Ts U T = {bound_Us = bound_Ts, bound_Ts = bound_Ts, U = U, T = T}; fun rebuild_function_after_exploration new_fn explore {bound_Ts, ...} arg_ts = let val binder_types_old_fn = map (curry fastype_of1 bound_Ts) arg_ts; val binder_types_new_fn = new_fn |> binder_types o (curry fastype_of1 bound_Ts) |> take (length binder_types_old_fn); val paramss = map2 (build_params bound_Ts) binder_types_new_fn binder_types_old_fn; in map2 explore paramss arg_ts |> curry list_comb new_fn end; fun massage_map_corec explore {bound_Ts, U, T, ...} t = let val explore' = explore ooo build_params in massage_nested_corec_call ctxt has_self_call explore' explore' bound_Ts U T t end; fun massage_comp explore params t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => explore params (betapply (f1, (betapplys (f2, args)))) | _ => explore params t); fun massage_fun explore (params as {bound_Us, bound_Ts, U, T}) t = if can dest_funT T then let val arg_T = domain_type T; val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t); in add_boundvar t |> explore {bound_Us = arg_T :: bound_Us, bound_Ts = arg_T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_T, t)) end else explore params t fun massage_let_if_case_corec explore {bound_Ts, U, T, ...} t = massage_let_if_case ctxt has_self_call (fn bound_Ts => explore (build_params bound_Ts U T)) (K (unexpected_corec_call_in ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t])) bound_Ts t; val massage_map_let_if_case = massage_star [massage_map_corec, massage_fun, massage_comp, massage_let_if_case_corec]; fun explore_arg _ t = if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a nested corecursive call" ^ (if could_be_friend then " (try specifying \"(friend)\")" else "")) else t; fun explore_inner params t = massage_map_let_if_case explore_inner_general params t and explore_inner_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in if has_self_call t then (case as_member_of (#friends inner_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer) | NONE => (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#ctr_wrapper inner_buffer) |> curry (op $) (#Oper inner_buffer) | NONE => if is_self_call f_t then if friend andalso exists has_self_call arg_ts then (case Symtab.lookup (#friends inner_buffer) fun_name of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer)) else let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (#VLeaf inner_buffer) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) else #CLeaf inner_buffer $ t end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_inner_general params t; fun explore_outer params t = massage_map_let_if_case explore_outer_general params t and explore_outer_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#VLeaf outer_buffer) | NONE => if not (has_self_call t) then t |> expand_to_ctr_term ctxt T |> massage_let_if_case_corec explore_outer_general params else (case as_member_of (#friends outer_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_outer params arg_ts |> curry (op $) (#Oper outer_buffer) | NONE => if is_self_call f_t then let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (inner_fp_of f_t) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_outer_general params t; in (args, rhs |> explore_outer (build_params [] outer_ssig_T res_T) |> abs_tuple_balanced args) end; fun mk_corec_fun_def_rhs ctxt arg_Ts corecUU0 corecUU_arg = let val corecUU = enforce_type ctxt domain_type (fastype_of corecUU_arg) corecUU0 in abs_curried_balanced arg_Ts (corecUU $ unfold_id_bnf_etc ctxt corecUU_arg) end; fun get_options ctxt opts = let val plugins = get_first (fn Plugins_Option f => SOME (f ctxt) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; val friend = exists (can (fn Friend_Option => ())) opts; val transfer = exists (can (fn Transfer_Option => ())) opts; in (plugins, friend, transfer) end; fun add_function binding parsed_eq lthy = let fun pat_completeness_auto ctxt = Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt; val ({defname, pelims = [[pelim]], pinducts = [pinduct], psimps = [psimp], ...}, lthy') = Function.add_function [(Binding.concealed binding, NONE, NoSyn)] [(((Binding.concealed Binding.empty, []), parsed_eq), [], [])] Function_Common.default_config pat_completeness_auto lthy; in ((defname, (pelim, pinduct, psimp)), lthy') end; fun build_corecUU_arg_and_goals prove_termin (Free (fun_base_name, _)) (arg_ts, explored_rhs) lthy = let val inner_fp_name0 = fun_base_name ^ inner_fp_suffix; val inner_fp_free = Free (inner_fp_name0, fastype_of explored_rhs); in if Term.exists_subterm (curry (op aconv) inner_fp_free) explored_rhs then let val arg = mk_tuple_balanced arg_ts; val inner_fp_eq = mk_Trueprop_eq (betapply (inner_fp_free, arg), betapply (explored_rhs, arg)); val ((inner_fp_name, (pelim, pinduct, psimp)), lthy') = add_function (Binding.name inner_fp_name0) inner_fp_eq lthy; fun mk_triple elim induct simp = ([elim], [induct], [simp]); fun prepare_termin () = let val {goal, ...} = Proof.goal (Function.termination NONE lthy'); val termin_goal = goal |> Thm.concl_of |> Logic.unprotect |> Envir.beta_eta_contract; in (lthy', (mk_triple pelim pinduct psimp, [termin_goal])) end; val (lthy'', (inner_fp_triple, termin_goals)) = if prove_termin then (case try (Function.prove_termination NONE (Function_Common.termination_prover_tac true lthy')) lthy' of NONE => prepare_termin () | SOME ({elims = SOME [[elim]], inducts = SOME [induct], simps = SOME [simp], ...}, lthy'') => (lthy'', (mk_triple elim induct simp, []))) else prepare_termin (); val inner_fp_const = (Binding.name_of inner_fp_name, fastype_of explored_rhs) |>> Proof_Context.read_const {proper = true, strict = false} lthy' |> (fn (Const (s, _), T) => Const (s, T)); in (((inner_fp_triple, termin_goals), inner_fp_const), lthy'') end else (((([], [], []), []), explored_rhs), lthy) end; fun derive_eq_corecUU ctxt {sig_fp_sugars, ssig_fp_sugar, eval, corecUU, eval_simps, all_algLam_algs, corecUU_unique, ...} fun_t corecUU_arg fun_code = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt (Thm.prop_of fun_code); val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val def_rhs = mk_corec_fun_def_rhs ctxt arg_Ts corecUU corecUU_arg; val goal = mk_Trueprop_eq (fun_t, def_rhs); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_corecUU_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique fun_code) |> Thm.close_derivation \<^here> end; fun derive_coinduct_cong_intros ({fpT = fpT0 as Type (fpT_name, _), friend_names = friend_names0, corecUU = Const (corecUU_name, _), dtor_coinduct_info as {dtor_coinduct, ...}, ...}) lthy = let val thy = Proof_Context.theory_of lthy; val phi = Proof_Context.export_morphism lthy (Local_Theory.target_of lthy); val fpT = Morphism.typ phi fpT0; val general_fpT = body_type (Sign.the_const_type thy corecUU_name); val most_general = Sign.typ_instance thy (general_fpT, fpT); in (case (most_general, coinduct_extra_of lthy corecUU_name) of (true, SOME extra) => ((false, extra), lthy) | _ => let val ctr_names = ctr_names_of_fp_name lthy fpT_name; val friend_names = friend_names0 |> map Long_Name.base_name |> rev; val cong_intro_pairs = derive_cong_intros lthy ctr_names friend_names dtor_coinduct_info; val (coinduct, coinduct_attrs) = derive_coinduct lthy fpT0 dtor_coinduct; val ((_, [coinduct]), lthy) = (* TODO check: only if most_general?*) Local_Theory.note ((Binding.empty, coinduct_attrs), [coinduct]) lthy; val extra = {coinduct = coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = cong_intro_pairs}; in ((most_general, extra), lthy |> most_general ? register_coinduct_extra corecUU_name extra) end) end; fun update_coinduct_cong_intross_dynamic fpT_name lthy = let val all_corec_infos = corec_infos_of lthy fpT_name in lthy |> fold_map (apfst snd oo derive_coinduct_cong_intros) all_corec_infos |> snd end; fun derive_and_update_coinduct_cong_intross [] = pair (false, []) | derive_and_update_coinduct_cong_intross (corec_infos as {fpT = Type (fpT_name, _), ...} :: _) = fold_map derive_coinduct_cong_intros corec_infos #>> split_list #> (fn ((changeds, extras), lthy) => if exists I changeds then ((true, extras), lthy |> update_coinduct_cong_intross_dynamic fpT_name) else ((false, extras), lthy)); fun prepare_corec_ursive_cmd int long_cmd opts (raw_fixes, raw_eq) lthy = let val _ = can the_single raw_fixes orelse error "Mutually corecursive functions not supported"; val (plugins, friend, transfer) = get_options lthy opts; val ([((b, fun_T), mx)], [(_, eq)]) = fst (Specification.read_multi_specs raw_fixes [((Binding.empty_atts, raw_eq), [], [])] lthy); val _ = check_top_sort lthy b fun_T; val (arg_Ts, res_T) = strip_type fun_T; val fpT_name = (case res_T of Type (s, _) => s | _ => not_codatatype lthy res_T); val fun_free = Free (Binding.name_of b, fun_T); val parsed_eq = parse_corec_equation lthy [fun_free] eq; val fun_name = Local_Theory.full_name lthy b; val fun_t = Const (fun_name, fun_T); (* FIXME: does this work with locales that fix variables? *) val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; fun extract_rho lthy' = let val lthy'' = lthy' |> Variable.declare_typ fun_T; val (prepared as (_, _, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, _, ssig_fp_sugar, buffer), lthy''') = prepare_friend_corec fun_name fun_T lthy''; val friend_parse_info = friend_parse_info_of lthy''' arg_Ts res_T buffer; val parsed_eq' = parsed_eq ||> subst_atomic [(fun_free, fun_t)]; in lthy''' |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq' |>> pair prepared end; val ((prepareds, (rho_datas, transfer_goal_datas)), lthy2) = if friend then extract_rho lthy1 |>> (apfst single ##> (apfst single #> apsnd single)) else (([], ([], [])), lthy1); val ((buffer, corec_infos), lthy3) = if friend then ((#13 (the_single prepareds), []), lthy2) else corec_info_of res_T lthy2 ||> no_base ? update_coinduct_cong_intross_dynamic fpT_name |>> (fn info as {buffer, ...} => (buffer, [info])); val corec_parse_info = corec_parse_info_of lthy3 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy3 true friend fun_name fun_free corec_parse_info res_T parsed_eq; val (((inner_fp_triple, termin_goals), corecUU_arg), lthy4) = build_corecUU_arg_and_goals (not long_cmd) fun_free explored_eq lthy3; fun def_fun (inner_fp_elims0, inner_fp_inducts0, inner_fp_simps0) const_transfers rho_transfers_foldeds lthy5 = let fun register_friend lthy' = let val [(old_corec_info, fp_b, version, Y, Z, _, k_T, _, _, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, _)] = prepareds; val [(rho, rho_def)] = rho_datas; val [(_, rho_transfer_goal)] = transfer_goal_datas; val Type (fpT_name, _) = res_T; val rho_transfer_folded = (case rho_transfers_foldeds of [] => derive_rho_transfer_folded lthy' fpT_name const_transfers rho_def rho_transfer_goal | [thm] => thm); in lthy' |> register_coinduct_dynamic_friend fpT_name fun_name |> register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho rho_transfer_folded old_corec_info end; val (friend_infos, lthy6) = lthy5 |> (if friend then register_friend #>> single else pair []); val (corec_info as {corecUU = corecUU0, ...}, lthy7) = (case corec_infos of [] => corec_info_of res_T lthy6 | [info] => (info, lthy6)); val def_rhs = mk_corec_fun_def_rhs lthy7 arg_Ts corecUU0 corecUU_arg; val def = ((b, mx), ((Binding.concealed (Thm.def_binding b), []), def_rhs)); val ((fun_lhs0, (_, fun_def0)), (lthy9, lthy8')) = lthy7 |> (snd o Local_Theory.begin_nested) |> Local_Theory.define def |> tap (fn (def, lthy') => print_def_consts int [def] lthy') ||> `Local_Theory.end_nested; val parsed_eq = parse_corec_equation lthy9 [fun_free] eq; val views0 = generate_views lthy9 eq fun_free parsed_eq; val lthy9' = lthy9 |> fold Variable.declare_typ (res_T :: arg_Ts); val phi = Proof_Context.export_morphism lthy8' lthy9'; val fun_lhs = Morphism.term phi fun_lhs0; val fun_def = Morphism.thm phi fun_def0; val inner_fp_elims = map (Morphism.thm phi) inner_fp_elims0; val inner_fp_inducts = map (Morphism.thm phi) inner_fp_inducts0; val inner_fp_simps = map (Morphism.thm phi) inner_fp_simps0; val (code_goal, _, _, _, _) = morph_views phi views0; fun derive_and_note_friend_extra_theorems lthy' = let val k_T = #7 (the_single prepareds); val rho_def = snd (the_single rho_datas); val (eq_algrho, algrho_eq) = derive_eq_algrho lthy' corec_info (the_single friend_infos) fun_lhs k_T code_goal const_transfers rho_def fun_def; val notes = (if Config.get lthy' bnf_internals then [(eq_algrhoN, [eq_algrho])] else []) |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false friendN (Binding.name thmN)), []), [(thms, [])])); in lthy' |> register_friend_extra fun_name eq_algrho algrho_eq |> Local_Theory.notes notes |> snd end; val lthy10 = lthy9 |> friend ? derive_and_note_friend_extra_theorems; val code_thm = derive_code lthy10 inner_fp_simps code_goal corec_info fun_lhs fun_def; (* TODO: val ctr_thmss = map mk_thm (#2 views); val disc_thmss = map mk_thm (#3 views); val disc_iff_thmss = map mk_thm (#4 views); val sel_thmss = map mk_thm (#5 views); *) val uniques = if null inner_fp_simps then [derive_unique lthy10 phi (#1 views0) corec_info fpT_name fun_def] else []; (* TODO: val disc_iff_or_disc_thmss = map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss; val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss; *) val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy11) = lthy10 |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val anonymous_notes = []; (* TODO: [(flat disc_iff_or_disc_thmss, simp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); *) val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (codeN, [code_thm], nitpicksimp_attrs), (coinductN, [coinduct], coinduct_attrs), (inner_inductN, inner_fp_inducts, []), (uniqueN, uniques, [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy11 bnf_internals then [(inner_elimN, inner_fp_elims, []), (inner_simpN, inner_fp_simps, [])] else []) (* TODO: (ctrN, ctr_thms, []), (discN, disc_thms, []), (disc_iffN, disc_iff_thms, []), (selN, sel_thms, simp_attrs), (simpsN, simp_thms, []), *) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false corecN (Binding.name thmN)), attrs), [(thms, [])])) |> filter_out (null o fst o hd o snd); in lthy11 (* TODO: |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat sel_thmss) |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat ctr_thmss) *) |> Spec_Rules.add Binding.empty Spec_Rules.equational [fun_lhs] [code_thm] |> plugins code_plugin ? Code.declare_default_eqns [(code_thm, true)] |> Local_Theory.notes (anonymous_notes @ notes) |> snd end; fun prove_transfer_goal ctxt goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (Transfer.transfer_prover_tac ctxt))) |> Thm.close_derivation \<^here>; fun maybe_prove_transfer_goal ctxt goal = (case try (prove_transfer_goal ctxt) goal of SOME thm => apfst (cons thm) | NONE => apsnd (cons goal)); val const_transfer_goals = fold (union (op aconv) o fst) transfer_goal_datas []; val (const_transfers, const_transfer_goals') = if long_cmd then ([], const_transfer_goals) else fold (maybe_prove_transfer_goal lthy4) const_transfer_goals ([], []); in ((def_fun, (([res_T], prepareds, rho_datas, map snd transfer_goal_datas), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals'))), lthy4) end; fun corec_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (_, (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int false opts (raw_fixes, raw_eq) lthy; in if not (null termin_goals) then error ("Termination prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else if not (null const_transfer_goals) then error ("Transfer prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else def_fun inner_fp_triple const_transfers [] lthy' end; fun corecursive_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (([Type (fpT_name, _)], prepareds, rho_datas, rho_transfer_goals), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int true opts (raw_fixes, raw_eq) lthy; val (rho_transfer_goals', unprime_rho_transfer_and_folds) = @{map 3} (fn (_, _, _, _, _, _, _, _, _, _, _, _, _) => fn (_, rho_def) => prime_rho_transfer_goal lthy' fpT_name rho_def) prepareds rho_datas rho_transfer_goals |> split_list; in Proof.theorem NONE (fn [termin_thms, const_transfers', rho_transfers'] => let val remove_domain_condition = full_simplify (put_simpset HOL_basic_ss lthy' addsimps (@{thm True_implies_equals} :: termin_thms)); in def_fun (@{apply 3} (map remove_domain_condition) inner_fp_triple) (const_transfers @ const_transfers') (map2 (fn f => f) unprime_rho_transfer_and_folds rho_transfers') end) (map (map (rpair [])) [termin_goals, const_transfer_goals, rho_transfer_goals']) lthy' end; fun friend_of_corec_cmd ((raw_fun_name, raw_fun_T_opt), raw_eq) lthy = let val Const (fun_name, _) = Proof_Context.read_const {proper = true, strict = false} lthy raw_fun_name; val fake_lthy = lthy |> (case raw_fun_T_opt of SOME raw_T => Proof_Context.add_const_constraint (fun_name, SOME (Syntax.read_typ lthy raw_T)) | NONE => I) handle TYPE (s, _, _) => error s; val fun_b = Binding.name (Long_Name.base_name fun_name); val code_goal = Syntax.read_prop fake_lthy raw_eq; val fun_T = (case code_goal of \<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _) => fastype_of (head_of t) | _ => ill_formed_equation_lhs_rhs lthy [code_goal]); val fun_t = Const (fun_name, fun_T); val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T; val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; val lthy2 = lthy1 |> Variable.declare_typ fun_T; val ((old_corec_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy3) = prepare_friend_corec fun_name fun_T lthy2; val friend_parse_info = friend_parse_info_of lthy3 arg_Ts res_T buffer; val parsed_eq = parse_corec_equation lthy3 [] code_goal; val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy4) = extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy3; fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info lthy5 = let val (corec_info, lthy6) = corec_info_of res_T lthy5; val fun_free = Free (Binding.name_of fun_b, fun_T); fun freeze_fun (t as Const (s, T)) = if s = fun_name andalso T = fun_T then fun_free else t | freeze_fun t = t; val eq = Term.map_aterms freeze_fun code_goal; val parsed_eq = parse_corec_equation lthy6 [fun_free] eq; val corec_parse_info = corec_parse_info_of lthy6 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy6 false false fun_name fun_free corec_parse_info res_T parsed_eq; val ((_, corecUU_arg), _) = build_corecUU_arg_and_goals false fun_free explored_eq lthy6; val eq_corecUU = derive_eq_corecUU lthy6 corec_info fun_t corecUU_arg code_thm; val (eq_algrho, algrho_eq) = derive_eq_algrho lthy6 corec_info friend_info fun_t k_T code_goal const_transfers rho_def eq_corecUU; val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy7) = lthy6 |> register_friend_extra fun_name eq_algrho algrho_eq |> register_coinduct_dynamic_friend fpT_name fun_name |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val unique = derive_unique lthy7 Morphism.identity code_goal corec_info fpT_name eq_corecUU; val notes = [(codeN, [code_thm], []), (coinductN, [coinduct], coinduct_attrs), (cong_introsN, maps snd cong_intros_pairs, []), (uniqueN, [unique], [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy7 bnf_internals then [(eq_algrhoN, [eq_algrho], []), (eq_corecUUN, [eq_corecUU], [])] else []) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of fun_b) (Binding.qualify false friendN (Binding.name thmN)), attrs), [(thms, [])])); in lthy7 |> Local_Theory.notes notes |> snd end; val (rho_transfer_goal', unprime_rho_transfer_and_fold) = prime_rho_transfer_goal lthy4 fpT_name rho_def rho_transfer_goal; in lthy4 |> Proof.theorem NONE (fn [[code_thm], const_transfers, [rho_transfer']] => register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho (unprime_rho_transfer_and_fold rho_transfer') old_corec_info #-> register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T) (map (map (rpair [])) [[code_goal], const_transfer_goals, [rho_transfer_goal']]) |> Proof.refine_singleton (Method.primitive_text (K I)) end; fun coinduction_upto_cmd (base_name, raw_fpT) lthy = let val fpT as Type (fpT_name, _) = Syntax.read_typ lthy raw_fpT; val no_base = has_no_corec_info lthy fpT_name; val (corec_info as {version, ...}, lthy1) = lthy |> corec_info_of fpT; val lthy2 = lthy1 |> no_base ? setup_base fpT_name; val ((changed, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy3) = lthy2 |> derive_and_update_coinduct_cong_intross [corec_info]; val lthy4 = lthy3 |> (changed orelse no_base) ? update_coinduct_cong_intross_dynamic fpT_name; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (coinduct_uptoN, [coinduct], coinduct_attrs)] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs |> map (fn (thmN, thms, attrs) => (((Binding.qualify true base_name (Binding.qualify false ("v" ^ string_of_int version) (Binding.name thmN))), attrs), [(thms, [])])); in lthy4 |> Local_Theory.notes notes |> snd end; fun consolidate lthy = let val corec_infoss = map (corec_infos_of lthy o fst) (all_codatatype_extras_of lthy); val (changeds, lthy') = lthy |> fold_map (apfst fst oo derive_and_update_coinduct_cong_intross) corec_infoss; in if exists I changeds then lthy' else raise Same.SAME end; fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Outer_Syntax.local_theory \<^command_keyword>\corec\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corec_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\corecursive\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corecursive_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\friend_of_corec\ "register a function as a legal context for nonprimitive corecursion" (Parse.const -- Scan.option (\<^keyword>\::\ |-- Parse.typ) --| Parse.where_ -- Parse.prop >> friend_of_corec_cmd); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinduction_upto\ "derive a coinduction up-to principle and a corresponding congruence closure" (Parse.name --| \<^keyword>\:\ -- Parse.typ >> coinduction_upto_cmd); val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_lfp_size.ML b/src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML @@ -1,399 +1,399 @@ (* Title: HOL/Tools/BNF/bnf_lfp_size.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2014 Generation of size functions for datatypes. *) signature BNF_LFP_SIZE = sig val register_size: string -> string -> thm -> thm list -> thm list -> local_theory -> local_theory val register_size_global: string -> string -> thm -> thm list -> thm list -> theory -> theory val size_of: Proof.context -> string -> (string * (thm * thm list * thm list)) option val size_of_global: theory -> string -> (string * (thm * thm list * thm list)) option end; structure BNF_LFP_Size : BNF_LFP_SIZE = struct open BNF_Util open BNF_Tactics open BNF_Def open BNF_FP_Def_Sugar val size_N = "size_"; val sizeN = "size"; val size_genN = "size_gen"; val size_gen_o_mapN = "size_gen_o_map"; val size_neqN = "size_neq"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun mk_plus_nat (t1, t2) = Const (\<^const_name>\Groups.plus\, HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2; fun mk_to_natT T = T --> HOLogic.natT; fun mk_abs_zero_nat T = Term.absdummy T HOLogic.zero; fun mk_unabs_def_unused_0 n = funpow n (fn thm => thm RS @{thm fun_cong_unused_0} handle THM _ => thm RS fun_cong); structure Data = Generic_Data ( type T = (string * (thm * thm list * thm list)) Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ); fun check_size_type thy T_name size_name = let val n = Sign.arity_number thy T_name; val As = map (fn s => TFree (s, \<^sort>\type\)) (Name.invent Name.context Name.aT n); val T = Type (T_name, As); val size_T = map mk_to_natT As ---> mk_to_natT T; val size_const = Const (size_name, size_T); in can (Thm.global_cterm_of thy) size_const orelse error ("Constant " ^ quote size_name ^ " registered as size function for " ^ quote T_name ^ " must have type\n" ^ quote (Syntax.string_of_typ_global thy size_T)) end; fun register_size T_name size_name overloaded_size_def size_simps size_gen_o_maps lthy = (check_size_type (Proof_Context.theory_of lthy) T_name size_name; Context.proof_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) lthy); fun register_size_global T_name size_name overloaded_size_def size_simps size_gen_o_maps thy = (check_size_type thy T_name size_name; Context.theory_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) thy); val size_of = Symtab.lookup o Data.get o Context.Proof; val size_of_global = Symtab.lookup o Data.get o Context.Theory; fun all_overloaded_size_defs_of ctxt = Symtab.fold (fn (_, (_, (overloaded_size_def, _, _))) => can (Logic.dest_equals o Thm.prop_of) overloaded_size_def ? cons overloaded_size_def) (Data.get (Context.Proof ctxt)) []; val size_gen_o_map_simps = @{thms inj_on_id snd_comp_apfst[simplified apfst_def]}; fun mk_size_gen_o_map_tac ctxt size_def rec_o_map inj_maps size_maps = unfold_thms_tac ctxt [size_def] THEN HEADGOAL (rtac ctxt (rec_o_map RS trans) THEN' asm_simp_tac (ss_only (inj_maps @ size_maps @ size_gen_o_map_simps) ctxt)) THEN IF_UNSOLVED (unfold_thms_tac ctxt @{thms id_def o_def} THEN HEADGOAL (rtac ctxt refl)); fun mk_size_neq ctxt cts exhaust sizes = HEADGOAL (rtac ctxt (infer_instantiate' ctxt (map SOME cts) exhaust)) THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' rtac ctxt @{thm trans_less_add2})); fun generate_datatype_size (fp_sugars as ({T = Type (_, As), BT = Type (_, Bs), fp = Least_FP, fp_res = {bnfs = fp_bnfs, ...}, fp_nesting_bnfs, live_nesting_bnfs, fp_co_induct_sugar = SOME _, ...} : fp_sugar) :: _) lthy0 = let val data = Data.get (Context.Proof lthy0); val Ts = map #T fp_sugars val T_names = map (fst o dest_Type) Ts; val nn = length Ts; val B_ify = Term.typ_subst_atomic (As ~~ Bs); val recs = map (#co_rec o the o #fp_co_induct_sugar) fp_sugars; val rec_thmss = map (#co_rec_thms o the o #fp_co_induct_sugar) fp_sugars; val rec_Ts as rec_T1 :: _ = map fastype_of recs; val rec_arg_Ts = binder_fun_types rec_T1; val Cs = map body_type rec_Ts; val Cs_rho = map (rpair HOLogic.natT) Cs; val substCnatT = Term.subst_atomic_types Cs_rho; val f_Ts = map mk_to_natT As; val f_TsB = map mk_to_natT Bs; val num_As = length As; fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names f_Ts; val fsB = map2 (curry Free) f_names f_TsB; val As_fs = As ~~ fs; val size_bs = map ((fn base => Binding.qualify false base (Binding.name (prefix size_N base))) o Long_Name.base_name) T_names; fun is_prod_C \<^type_name>\prod\ [_, T'] = member (op =) Cs T' | is_prod_C _ _ = false; fun mk_size_of_typ (T as TFree _) = pair (case AList.lookup (op =) As_fs T of SOME f => f | NONE => if member (op =) Cs T then Term.absdummy T (Bound 0) else mk_abs_zero_nat T) | mk_size_of_typ (T as Type (s, Ts)) = if is_prod_C s Ts then pair (snd_const T) else if exists (exists_subtype_in (As @ Cs)) Ts then (case Symtab.lookup data s of SOME (size_name, (_, _, size_gen_o_maps)) => let val (args, size_gen_o_mapss') = fold_map mk_size_of_typ Ts []; val size_T = map fastype_of args ---> mk_to_natT T; val size_const = Const (size_name, size_T); in append (size_gen_o_maps :: size_gen_o_mapss') #> pair (Term.list_comb (size_const, args)) end | _ => pair (mk_abs_zero_nat T)) else pair (mk_abs_zero_nat T); fun mk_size_of_arg t = mk_size_of_typ (fastype_of t) #>> (fn s => substCnatT (betapply (s, t))); fun is_recursive_or_plain_case Ts = exists (exists_subtype_in Cs) Ts orelse forall (not o exists_subtype_in As) Ts; (* We want the size function to enjoy the following properties: 1. The size of a list should coincide with its length. 2. All the nonrecursive constructors of a type should have the same size. 3. Each constructor through which nested recursion takes place should count as at least one in the generic size function. 4. The "size" function should be definable as "size_t (%_. 0) ... (%_. 0)", where "size_t" is the generic function. This explains the somewhat convoluted logic ahead. *) val base_case = if forall (is_recursive_or_plain_case o binder_types) rec_arg_Ts then HOLogic.zero else HOLogic.Suc_zero; fun mk_size_arg rec_arg_T = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val (summands, size_gen_o_mapss) = fold_map mk_size_of_arg xs [] |>> remove (op =) HOLogic.zero; val sum = if null summands then base_case else foldl1 mk_plus_nat (summands @ [HOLogic.Suc_zero]); in append size_gen_o_mapss #> pair (fold_rev Term.lambda (map substCnatT xs) sum) end; fun mk_size_rhs recx = fold_map mk_size_arg rec_arg_Ts #>> (fn args => fold_rev Term.lambda fs (Term.list_comb (substCnatT recx, args))); val maybe_conceal_def_binding = Thm.def_binding #> not (Config.get lthy0 bnf_internals) ? Binding.concealed; val (size_rhss, nested_size_gen_o_mapss) = fold_map mk_size_rhs recs []; val size_Ts = map fastype_of size_rhss; val nested_size_gen_o_maps_complete = forall (not o null) nested_size_gen_o_mapss; val nested_size_gen_o_maps = fold (union Thm.eq_thm_prop) nested_size_gen_o_mapss []; val ((raw_size_consts, raw_size_defs), (lthy1, lthy1_old)) = lthy0 |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 2} (fn b => fn rhs => Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd) size_bs size_rhss ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy1_old lthy1; val size_defs = map (Morphism.thm phi) raw_size_defs; val size_consts0 = map (Morphism.term phi) raw_size_consts; val size_consts = map2 retype_const_or_free size_Ts size_consts0; val size_constsB = map (Term.map_types B_ify) size_consts; val zeros = map mk_abs_zero_nat As; val overloaded_size_rhss = map (fn c => Term.list_comb (c, zeros)) size_consts; val overloaded_size_Ts = map fastype_of overloaded_size_rhss; val overloaded_size_consts = map (curry Const \<^const_name>\size\) overloaded_size_Ts; val overloaded_size_def_bs = map (maybe_conceal_def_binding o Binding.suffix_name "_overloaded") size_bs; fun define_overloaded_size def_b lhs0 rhs lthy = let val Free (c, _) = Syntax.check_term lthy lhs0; val ((_, (_, thm)), lthy') = lthy |> Local_Theory.define ((Binding.name c, NoSyn), ((def_b, []), rhs)); val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy'); val thm' = singleton (Proof_Context.export lthy' thy_ctxt) thm; in (thm', lthy') end; val (overloaded_size_defs, lthy2) = lthy1 |> Local_Theory.background_theory_result (Class.instantiation (T_names, map dest_TFree As, [HOLogic.class_size]) #> @{fold_map 3} define_overloaded_size overloaded_size_def_bs overloaded_size_consts overloaded_size_rhss ##> Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt []) ##> Local_Theory.exit_global); val size_defs' = map (mk_unabs_def (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val size_defs_unused_0 = map (mk_unabs_def_unused_0 (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val overloaded_size_defs' = map (mk_unabs_def 1 o HOLogic.mk_obj_eq) overloaded_size_defs; val nested_size_maps = map (mk_pointful lthy2) nested_size_gen_o_maps @ nested_size_gen_o_maps; val all_inj_maps = @{thm prod.inj_map} :: map inj_map_of_bnf (fp_bnfs @ fp_nesting_bnfs @ live_nesting_bnfs) |> distinct Thm.eq_thm_prop; fun derive_size_simp size_def' simp0 = (trans OF [size_def', simp0]) |> Simplifier.asm_full_simplify (ss_only (@{thms inj_on_convol_ident id_def o_def snd_conv} @ all_inj_maps @ nested_size_maps) lthy2) |> Local_Defs.fold lthy2 size_defs_unused_0; fun derive_overloaded_size_simp overloaded_size_def' simp0 = (trans OF [overloaded_size_def', simp0]) |> unfold_thms lthy2 @{thms add_0_left add_0_right} |> Local_Defs.fold lthy2 (overloaded_size_defs @ all_overloaded_size_defs_of lthy2); val size_simpss = map2 (map o derive_size_simp) size_defs' rec_thmss; val size_simps = flat size_simpss; val overloaded_size_simpss = map2 (map o derive_overloaded_size_simp) overloaded_size_defs' size_simpss; val overloaded_size_simps = flat overloaded_size_simpss; val size_thmss = map2 append size_simpss overloaded_size_simpss; val size_gen_thmss = size_simpss; fun rhs_is_zero thm = let val Const (trueprop, _) $ (Const (eq, _) $ _ $ rhs) = Thm.prop_of thm in trueprop = \<^const_name>\Trueprop\ andalso eq = \<^const_name>\HOL.eq\ andalso rhs = HOLogic.zero end; val size_neq_thmss = @{map 3} (fn fp_sugar => fn size => fn size_thms => if exists rhs_is_zero size_thms then [] else let val (xs, _) = mk_Frees "x" (binder_types (fastype_of size)) lthy2; val goal = HOLogic.mk_Trueprop (BNF_LFP_Util.mk_not_eq (list_comb (size, xs)) HOLogic.zero); val vars = Variable.add_free_names lthy2 goal []; val thm = Goal.prove_sorry lthy2 vars [] goal (fn {context = ctxt, ...} => mk_size_neq ctxt (map (Thm.cterm_of lthy2) xs) (#exhaust (#ctr_sugar (#fp_ctr_sugar fp_sugar))) size_thms) |> single |> map (Thm.close_derivation \<^here>); in thm end) fp_sugars overloaded_size_consts overloaded_size_simpss; val ABs = As ~~ Bs; val g_names = variant_names num_As "g"; val gs = map2 (curry Free) g_names (map (op -->) ABs); val liveness = map (op <>) ABs; val live_gs = AList.find (op =) (gs ~~ liveness) true; val live = length live_gs; val maps0 = map map_of_bnf fp_bnfs; val size_gen_o_map_thmss = if live = 0 then replicate nn [] else let val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val size_gen_o_map_conds = if exists (can Logic.dest_implies o Thm.prop_of) nested_size_gen_o_maps then map (HOLogic.mk_Trueprop o mk_inj) live_gs else []; val fsizes = map (fn size_constB => Term.list_comb (size_constB, fsB)) size_constsB; val size_gen_o_map_lhss = map2 (curry HOLogic.mk_comp) fsizes gmaps; val fgs = map2 (fn fB => fn g as Free (_, Type (_, [A, B])) => if A = B then fB else HOLogic.mk_comp (fB, g)) fsB gs; val size_gen_o_map_rhss = map (fn c => Term.list_comb (c, fgs)) size_consts; val size_gen_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fsB, gs] o curry Logic.list_implies size_gen_o_map_conds o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) size_gen_o_map_lhss size_gen_o_map_rhss; val rec_o_maps = fold_rev (curry (op @) o #co_rec_o_maps o the o #fp_co_induct_sugar) fp_sugars []; val size_gen_o_map_thmss = if nested_size_gen_o_maps_complete andalso forall (fn TFree (_, S) => S = \<^sort>\type\) As then @{map 3} (fn goal => fn size_def => fn rec_o_map => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_size_gen_o_map_tac ctxt size_def rec_o_map all_inj_maps nested_size_maps) |> Thm.close_derivation \<^here> |> single) size_gen_o_map_goals size_defs rec_o_maps else replicate nn []; in size_gen_o_map_thmss end; val massage_multi_notes = maps (fn (thmN, thmss, attrs) => map2 (fn T_name => fn thms => ((Binding.qualify true (Long_Name.base_name T_name) (Binding.name thmN), attrs), [(thms, [])])) T_names thmss) #> filter_out (null o fst o hd o snd); val notes = [(sizeN, size_thmss, nitpicksimp_attrs @ simp_attrs), (size_genN, size_gen_thmss, []), (size_gen_o_mapN, size_gen_o_map_thmss, []), (size_neqN, size_neq_thmss, [])] |> massage_multi_notes; val (noted, lthy3) = lthy2 |> Spec_Rules.add Binding.empty Spec_Rules.equational size_consts size_simps |> Spec_Rules.add Binding.empty Spec_Rules.equational overloaded_size_consts overloaded_size_simps |> Code.declare_default_eqns (map (rpair true) (flat size_thmss)) (*Ideally, this would be issued only if the "code" plugin is enabled.*) |> Local_Theory.notes notes; val phi0 = substitute_noted_thm noted; in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (@{fold 3} (fn T_name => fn Const (size_name, _) => fn overloaded_size_def => let val morph = Morphism.thm (phi0 $> phi) in Symtab.update (T_name, (size_name, (morph overloaded_size_def, map morph overloaded_size_simps, maps (map morph) size_gen_o_map_thmss))) end) T_names size_consts overloaded_size_defs)) end | generate_datatype_size _ lthy = lthy; val size_plugin = Plugin_Name.declare_setup \<^binding>\size\; val _ = Theory.setup (fp_sugars_interpretation size_plugin generate_datatype_size); end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1272 +1,1272 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => let val pos = Position.thread_data () in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = let val tab = Data.get (Context.Theory thy); val pos = Position.thread_data (); in if Symtab.defined tab name then thy else thy |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^Const_>\Not for \t' $ Bound 0\\) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, \<^Const_>\HOL.eq _ for \Bound 0\ t'\) => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^Const_>\Not for \<^Const_>\HOL.eq _ for \Bound 0\ t'\\) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (TVars.make rho_As, Vars.empty) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') (Thm.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^Const>\True\ else \<^Const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; - val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); - val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); + val exhaust_case_names_attr = Attrib.internal \<^here> (K (Rule_Cases.case_names exhaust_cases)); + val cases_type_attr = Attrib.internal \<^here> (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) - #> Local_Theory.declaration {syntax = false, pervasive = false} + #> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); (** external views **) (* document antiquotations *) local fun antiquote_setup binding co = Document_Output.antiquotation_pretty_source_embedded binding ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- Args.type_name {proper = true, strict = true}) (fn ctxt => fn (pos, type_name) => let fun err () = error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); in (case ctr_sugar_of ctxt type_name of NONE => err () | SOME {kind, T = T0, ctrs = ctrs0, ...} => let val _ = if co = (kind = Codatatype) then () else err (); val T = Logic.unvarifyT_global T0; val ctrs = map Logic.unvarify_global ctrs0; val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); fun pretty_ctr ctr = Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: map pretty_typ_bracket (binder_types (fastype_of ctr)))); in Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) end) end); in val _ = Theory.setup (antiquote_setup \<^binding>\datatype\ false #> antiquote_setup \<^binding>\codatatype\ true); end; (* theory export *) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); val datatypes = (Data.get (Context.Theory thy), []) |-> Symtab.fold (fn (name, (pos, {kind, T, ctrs, ...})) => if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I else let val pos_properties = Thy_Info.adjust_pos_properties context pos; val typ = Logic.unvarifyT_global T; val constrs = map Logic.unvarify_global ctrs; val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); val constructors = map (fn t => (t, Term.type_of t)) constrs; in cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) end); in if null datatypes then () else Export_Theory.export_body thy "datatypes" let open XML.Encode Term_XML.Encode in list (pair properties (pair string (pair bool (pair (list (pair string sort)) (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes end end else ()); end; diff --git a/src/HOL/Tools/Function/function.ML b/src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML +++ b/src/HOL/Tools/Function/function.ML @@ -1,288 +1,288 @@ (* Title: HOL/Tools/Function/function.ML Author: Alexander Krauss, TU Muenchen Main entry points to the function package. *) signature FUNCTION = sig type info = Function_Common.info val add_function: (binding * typ option * mixfix) list -> Specification.multi_specs -> Function_Common.function_config -> (Proof.context -> tactic) -> local_theory -> info * local_theory val add_function_cmd: (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> Function_Common.function_config -> (Proof.context -> tactic) -> bool -> local_theory -> info * local_theory val function: (binding * typ option * mixfix) list -> Specification.multi_specs -> Function_Common.function_config -> local_theory -> Proof.state val function_cmd: (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> Function_Common.function_config -> bool -> local_theory -> Proof.state val prove_termination: term option -> tactic -> local_theory -> info * local_theory val prove_termination_cmd: string option -> tactic -> local_theory -> info * local_theory val termination : term option -> local_theory -> Proof.state val termination_cmd : string option -> local_theory -> Proof.state val get_congs : Proof.context -> thm list val get_info : Proof.context -> term -> info end structure Function : FUNCTION = struct open Function_Lib open Function_Common val simp_attribs = @{attributes [simp, nitpick_simp]} val psimp_attribs = @{attributes [nitpick_psimp]} fun note_derived (a, atts) (fname, thms) = Local_Theory.note ((derived_name fname a, atts), thms) #> apfst snd fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy = let val spec = post simps |> map (apfst (apsnd (fn ats => moreatts @ ats))) |> map (apfst (apfst extra_qualify)) val (saved_spec_simps, lthy') = fold_map Local_Theory.note spec lthy val saved_simps = maps snd saved_spec_simps val simps_by_f = sort saved_simps fun note fname simps = Local_Theory.note ((mod_binding (derived_name fname label), []), simps) #> snd in (saved_simps, fold2 note fnames simps_by_f lthy') end fun prepare_function do_print prep fixspec eqns config lthy = let val ((fixes0, spec0), ctxt') = prep fixspec eqns lthy val fixes = map (apfst (apfst Binding.name_of)) fixes0 val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0 val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec val fnames = map (fst o fst) fixes0 val defname = Binding.conglomerate fnames; val FunctionConfig {partials, default, ...} = config val _ = if is_some default then legacy_feature "\"function (default)\" -- use 'partial_function' instead" else () val ((goal_state, cont), lthy') = Function_Mutual.prepare_function_mutual config defname fixes0 eqs lthy fun afterqed [[proof]] lthy1 = let val result = cont lthy1 (Thm.close_derivation \<^here> proof) val FunctionResult {fs, R, dom, psimps, simple_pinducts, termination, domintros, cases, ...} = result val pelims = Function_Elims.mk_partial_elim_rules lthy1 result val concealed_partial = if partials then I else Binding.concealed val addsmps = add_simps fnames post sort_cont val (((((psimps', [pinducts']), [termination']), cases'), pelims'), lthy2) = lthy1 |> addsmps (concealed_partial o Binding.qualify false "partial") "psimps" concealed_partial psimp_attribs psimps ||>> Local_Theory.notes [((concealed_partial (derived_name defname "pinduct"), []), simple_pinducts |> map (fn th => ([th], [Attrib.case_names cnames, Attrib.consumes (1 - Thm.nprems_of th)] @ @{attributes [induct pred]})))] ||>> (apfst snd o Local_Theory.note ((Binding.concealed (derived_name defname "termination"), []), [termination])) ||>> fold_map (note_derived ("cases", [Attrib.case_names cnames])) (fnames ~~ map single cases) ||>> fold_map (note_derived ("pelims", [Attrib.consumes 1, Attrib.constraints 1])) (fnames ~~ pelims) ||> (case domintros of NONE => I | SOME thms => Local_Theory.note ((derived_name defname "domintros", []), thms) #> snd) val info = { add_simps=addsmps, fnames=fnames, case_names=cnames, psimps=psimps', pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination', totality=NONE, fs=fs, R=R, dom=dom, defname=defname, is_partial=true, cases=flat cases', pelims=pelims',elims=NONE} val _ = Proof_Display.print_consts do_print (Position.thread_data ()) lthy2 (K false) (map fst fixes) in (info, - lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false} + lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => add_function_data (transform_function_data phi info))) end in ((goal_state, afterqed), lthy') end fun gen_add_function do_print prep fixspec eqns config tac lthy = let val ((goal_state, afterqed), lthy') = prepare_function do_print prep fixspec eqns config lthy val pattern_thm = case SINGLE (tac lthy') goal_state of NONE => error "pattern completeness and compatibility proof failed" | SOME st => Goal.finish lthy' st in lthy' |> afterqed [[pattern_thm]] end val add_function = gen_add_function false Specification.check_multi_specs fun add_function_cmd a b c d int = gen_add_function int Specification.read_multi_specs a b c d fun gen_function do_print prep fixspec eqns config lthy = let val ((goal_state, afterqed), lthy') = prepare_function do_print prep fixspec eqns config lthy in lthy' |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]] |> Proof.refine_singleton (Method.primitive_text (K (K goal_state))) end val function = gen_function false Specification.check_multi_specs fun function_cmd a b c int = gen_function int Specification.read_multi_specs a b c fun prepare_termination_proof prep_binding prep_term raw_term_opt lthy = let val term_opt = Option.map (prep_term lthy) raw_term_opt val info = (case term_opt of SOME t => (case import_function_data t lthy of SOME info => info | NONE => error ("Not a function: " ^ quote (Syntax.string_of_term lthy t))) | NONE => (case import_last_function lthy of SOME info => info | NONE => error "Not a function")) val { termination, fs, R, add_simps, case_names, psimps, pinducts, defname, fnames, cases, dom, pelims, ...} = info val domT = domain_type (fastype_of R) val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT))) fun afterqed [[raw_totality]] lthy1 = let val totality = Thm.close_derivation \<^here> raw_totality val remove_domain_condition = full_simplify (put_simpset HOL_basic_ss lthy1 addsimps [totality, @{thm True_implies_equals}]) val tsimps = map remove_domain_condition psimps val tinduct = map remove_domain_condition pinducts val telims = map (map remove_domain_condition) pelims in lthy1 |> add_simps prep_binding "simps" prep_binding simp_attribs tsimps ||> Code.declare_default_eqns (map (rpair true) tsimps) ||>> Local_Theory.note ((prep_binding (derived_name defname "induct"), [Attrib.case_names case_names]), tinduct) ||>> fold_map (note_derived ("elims", [Attrib.consumes 1, Attrib.constraints 1])) (map prep_binding fnames ~~ telims) |-> (fn ((simps,(_,inducts)), elims) => fn lthy2 => let val info' = { is_partial=false, defname=defname, fnames=fnames, add_simps=add_simps, case_names=case_names, fs=fs, R=R, dom=dom, psimps=psimps, pinducts=pinducts, simps=SOME simps, inducts=SOME inducts, termination=termination, totality=SOME totality, cases=cases, pelims=pelims, elims=SOME elims} |> transform_function_data (Morphism.binding_morphism "" prep_binding) in (info', lthy2 - |> Local_Theory.declaration {syntax = false, pervasive = false} + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => add_function_data (transform_function_data phi info')) |> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef fs tsimps) end) end in (goal, afterqed, termination) end fun gen_prove_termination prep_term raw_term_opt tac lthy = let val (goal, afterqed, termination) = prepare_termination_proof I prep_term raw_term_opt lthy val totality = Goal.prove lthy [] [] goal (K tac) in afterqed [[totality]] lthy end val prove_termination = gen_prove_termination Syntax.check_term val prove_termination_cmd = gen_prove_termination Syntax.read_term fun gen_termination prep_term raw_term_opt lthy = let val (goal, afterqed, termination) = prepare_termination_proof Binding.reset_pos prep_term raw_term_opt lthy in lthy |> Proof_Context.note_thms "" ((Binding.empty, [Context_Rules.rule_del]), [([allI], [])]) |> snd |> Proof_Context.note_thms "" ((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])]) |> snd |> Proof_Context.note_thms "" ((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]), [([Goal.norm_result lthy termination], [])]) |> snd |> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]] end val termination = gen_termination Syntax.check_term val termination_cmd = gen_termination Syntax.read_term (* Datatype hook to declare datatype congs as "function_congs" *) fun add_case_cong n thy = let val cong = #case_cong (Old_Datatype_Data.the_info thy n) |> safe_mk_meta_eq in Context.theory_map (Function_Context_Tree.add_function_cong cong) thy end val _ = Theory.setup (Old_Datatype_Data.interpretation (K (fold add_case_cong))) (* get info *) val get_congs = Function_Context_Tree.get_function_congs fun get_info ctxt t = Function_Common.retrieve_function_data ctxt t |> the_single |> snd (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof' \<^command_keyword>\function\ "define general recursive functions" (function_parser default_config >> (fn (config, (fixes, specs)) => function_cmd fixes specs config)) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\termination\ "prove termination of a recursive function" (Scan.option Parse.term >> termination_cmd) end diff --git a/src/HOL/Tools/Lifting/lifting_def.ML b/src/HOL/Tools/Lifting/lifting_def.ML --- a/src/HOL/Tools/Lifting/lifting_def.ML +++ b/src/HOL/Tools/Lifting/lifting_def.ML @@ -1,675 +1,675 @@ (* Title: HOL/Tools/Lifting/lifting_def.ML Author: Ondrej Kuncar Definitions for constants on quotient types. *) signature LIFTING_DEF = sig datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ type lift_def val rty_of_lift_def: lift_def -> typ val qty_of_lift_def: lift_def -> typ val rhs_of_lift_def: lift_def -> term val lift_const_of_lift_def: lift_def -> term val def_thm_of_lift_def: lift_def -> thm val rsp_thm_of_lift_def: lift_def -> thm val abs_eq_of_lift_def: lift_def -> thm val rep_eq_of_lift_def: lift_def -> thm option val code_eq_of_lift_def: lift_def -> code_eq val transfer_rules_of_lift_def: lift_def -> thm list val morph_lift_def: morphism -> lift_def -> lift_def val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def val mk_lift_const_of_lift_def: typ -> lift_def -> term type config = { notes: bool } val map_config: (bool -> bool) -> config -> config val default_config: config val generate_parametric_transfer_rule: Proof.context -> thm -> thm -> thm val add_lift_def: config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory val prepare_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> thm list -> local_theory -> term option * (thm -> Proof.context -> lift_def * local_theory) val gen_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val lift_def: config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val can_generate_code_cert: thm -> bool end structure Lifting_Def: LIFTING_DEF = struct open Lifting_Util infix 0 MRSL datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ datatype lift_def = LIFT_DEF of { rty: typ, qty: typ, rhs: term, lift_const: term, def_thm: thm, rsp_thm: thm, abs_eq: thm, rep_eq: thm option, code_eq: code_eq, transfer_rules: thm list }; fun rep_lift_def (LIFT_DEF lift_def) = lift_def; val rty_of_lift_def = #rty o rep_lift_def; val qty_of_lift_def = #qty o rep_lift_def; val rhs_of_lift_def = #rhs o rep_lift_def; val lift_const_of_lift_def = #lift_const o rep_lift_def; val def_thm_of_lift_def = #def_thm o rep_lift_def; val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def; val abs_eq_of_lift_def = #abs_eq o rep_lift_def; val rep_eq_of_lift_def = #rep_eq o rep_lift_def; val code_eq_of_lift_def = #code_eq o rep_lift_def; val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def; fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules = LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }; fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }) = LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const, def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq, code_eq = f9 code_eq, transfer_rules = f10 transfer_rules } fun morph_lift_def phi = let val mtyp = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm) end fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty) fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def) (lift_const_of_lift_def lift_def) fun inst_of_lift_def ctxt qty lift_def = let val instT = Vartab.fold (fn (a, (S, T)) => cons ((a, S), Thm.ctyp_of ctxt T)) (mk_inst_of_lift_def qty lift_def) [] val phi = Morphism.instantiate_morphism (TVars.make instT, Vars.empty) in morph_lift_def phi lift_def end (* Config *) type config = { notes: bool }; fun map_config f1 { notes = notes } = { notes = f1 notes } val default_config = { notes = true }; (* Reflexivity prover *) fun mono_eq_prover ctxt prop = let val refl_rules = Lifting_Info.get_reflexivity_rules ctxt val transfer_rules = Transfer.get_transfer_raw ctxt fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i in SOME (Goal.prove ctxt [] [] prop (K (main_tac 1))) handle ERROR _ => NONE end fun try_prove_refl_rel ctxt rel = let fun mk_ge_eq x = let val T = fastype_of x in Const (\<^const_name>\less_eq\, T --> T --> HOLogic.boolT) $ (Const (\<^const_name>\HOL.eq\, T)) $ x end; val goal = HOLogic.mk_Trueprop (mk_ge_eq rel); in mono_eq_prover ctxt goal end; fun try_prove_reflexivity ctxt prop = let val cprop = Thm.cterm_of ctxt prop val rule = @{thm ge_eq_refl} val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule) val insts = Thm.first_order_match (concl_pat, cprop) val rule = Drule.instantiate_normalize insts rule val prop = hd (Thm.prems_of rule) in case mono_eq_prover ctxt prop of SOME thm => SOME (thm RS rule) | NONE => NONE end (* Generates a parametrized transfer rule. transfer_rule - of the form T t f parametric_transfer_rule - of the form par_R t' t Result: par_T t' f, after substituing op= for relations in par_R that relate a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f using Lifting_Term.merge_transfer_relations *) fun generate_parametric_transfer_rule ctxt0 transfer_rule parametric_transfer_rule = let fun preprocess ctxt thm = let val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm; val param_rel = (snd o dest_comb o fst o dest_comb) tm; val free_vars = Term.add_vars param_rel []; fun make_subst (xi, typ) subst = let val [rty, rty'] = binder_types typ in if Term.is_TVar rty andalso Term.is_Type rty' then (xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst else subst end; val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm; in Conv.fconv_rule ((Conv.concl_conv (Thm.nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv) (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm end fun inst_relcomppI ctxt ant1 ant2 = let val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1 val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2 val fun1 = Thm.cterm_of ctxt (strip_args 2 t1) val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1) val fun2 = Thm.cterm_of ctxt (strip_args 2 t2) val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2) val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI} val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) [])) in infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI end fun zip_transfer_rules ctxt thm = let fun mk_POS ty = Const (\<^const_name>\POS\, ty --> ty --> HOLogic.boolT) val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm val typ = Thm.typ_of_cterm rel val POS_const = Thm.cterm_of ctxt (mk_POS typ) val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) val goal = Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var) in [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply} end val thm = inst_relcomppI ctxt0 parametric_transfer_rule transfer_rule OF [parametric_transfer_rule, transfer_rule] val preprocessed_thm = preprocess ctxt0 thm val (fixed_thm, ctxt1) = ctxt0 |> yield_singleton (apfst snd oo Variable.import true) preprocessed_thm val assms = cprems_of fixed_thm val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add val (prems, ctxt2) = ctxt1 |> fold_map Thm.assume_hyps assms val ctxt3 = ctxt2 |> Context.proof_map (fold add_transfer_rule prems) val zipped_thm = fixed_thm |> undisch_all |> zip_transfer_rules ctxt3 |> implies_intr_list assms |> singleton (Variable.export ctxt3 ctxt0) in zipped_thm end fun print_generate_transfer_info msg = let val error_msg = cat_lines ["Generation of a parametric transfer rule failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in error error_msg end fun map_ter _ x [] = x | map_ter f _ xs = map f xs fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms = let val transfer_rule = ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer}) |> Lifting_Term.parametrize_transfer_rule lthy in (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule])) end (* Generation of the code certificate from the rsp theorem *) fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V) | get_body_types (U, V) = (U, V) fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W) | get_binder_types _ = [] fun get_binder_types_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types_by_rel S (U, W) | get_binder_types_by_rel _ _ = [] fun get_body_type_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_type_by_rel S (U, V) | get_body_type_by_rel _ (U, V) = (U, V) fun get_binder_rels (Const (\<^const_name>\rel_fun\, _) $ R $ S) = R :: get_binder_rels S | get_binder_rels _ = [] fun force_rty_type ctxt rty rhs = let val thy = Proof_Context.theory_of ctxt val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs val rty_schematic = fastype_of rhs_schematic val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty in Envir.subst_term_types match rhs_schematic end fun unabs_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) fun dest_abs (Abs (var_name, T, _)) = (var_name, T) | dest_abs tm = raise TERM("get_abs_var",[tm]) val (var_name, T) = dest_abs (Thm.term_of rhs) val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T))) in Thm.combination def refl_thm |> singleton (Proof_Context.export ctxt' ctxt) end fun unabs_all_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) val xs = strip_abs_vars (Thm.term_of rhs) in fold (K (unabs_def ctxt)) xs def end val map_fun_unfolded = @{thm map_fun_def[abs_def]} |> unabs_def \<^context> |> unabs_def \<^context> |> Local_Defs.unfold0 \<^context> [@{thm comp_def}] fun unfold_fun_maps ctm = let fun unfold_conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\map_fun\, _) $ _ $ _ => (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm | _ => Conv.all_conv ctm in (Conv.fun_conv unfold_conv) ctm end fun unfold_fun_maps_beta ctm = let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false) in (unfold_fun_maps then_conv try_beta_conv) ctm end fun prove_rel ctxt rsp_thm (rty, qty) = let val ty_args = get_binder_types (rty, qty) fun disch_arg args_ty thm = let val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty in [quot_thm, thm] MRSL @{thm apply_rsp''} end in fold disch_arg ty_args rsp_thm end exception CODE_CERT_GEN of string fun simplify_code_eq ctxt def_thm = Local_Defs.unfold0 ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm (* quot_thm - quotient theorem (Quotient R Abs Rep T). returns: whether the Lifting package is capable to generate code for the abstract type represented by quot_thm *) fun can_generate_code_cert quot_thm = case quot_thm_rel quot_thm of Const (\<^const_name>\HOL.eq\, _) => true | Const (\<^const_name>\eq_onp\, _) $ _ => true | _ => false fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) = let val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm val unabs_def = unabs_all_def ctxt unfolded_def in if body_type rty = body_type qty then SOME (simplify_code_eq ctxt (HOLogic.mk_obj_eq unabs_def)) else let val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty)) val rel_fun = prove_rel ctxt rsp_thm (rty, qty) val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq} in case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of SOME mono_eq_thm => let val rep_abs_eq = mono_eq_thm RS rep_abs_thm val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm) val rep_refl = HOLogic.mk_obj_eq (Thm.reflexive rep) val repped_eq = [rep_refl, HOLogic.mk_obj_eq unabs_def] MRSL @{thm cong} val code_cert = [repped_eq, rep_abs_eq] MRSL trans in SOME (simplify_code_eq ctxt code_cert) end | NONE => NONE end end fun generate_abs_eq ctxt def_thm rsp_thm quot_thm = let val abs_eq_with_assms = let val (rty, qty) = quot_thm_rty_qty quot_thm val rel = quot_thm_rel quot_thm val ty_args = get_binder_types_by_rel rel (rty, qty) val body_type = get_body_type_by_rel rel (rty, qty) val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type val rep_abs_folded_unmapped_thm = let val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq} val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id) val unfolded_maps_eq = unfold_fun_maps ctm val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap} val prems_pat = (hd o Drule.cprems_of) t1 val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq) in unfolded_maps_eq RS (Drule.instantiate_normalize insts t1) end in rep_abs_folded_unmapped_thm |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm])) end val prem_rels = get_binder_rels (quot_thm_rel quot_thm); val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt) |> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the) |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl})) val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms in simplify_code_eq ctxt abs_eq end fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy = let fun no_abstr (t $ u) = no_abstr t andalso no_abstr u | no_abstr (Abs (_, _, t)) = no_abstr t | no_abstr (Const (name, _)) = not (Code.is_abstr thy name) | no_abstr _ = true fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) andalso no_abstr (Thm.prop_of eqn) fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq) in if is_valid_eq abs_eq_thm then (ABS_EQ, Code.declare_default_eqns_global [(abs_eq_thm, true)] thy) else let val (rty_body, qty_body) = get_body_types (rty, qty) in if rty_body = qty_body then (REP_EQ, Code.declare_default_eqns_global [(the opt_rep_eq_thm, true)] thy) else if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm) then (REP_EQ, Code.declare_abstract_eqn_global (the opt_rep_eq_thm) thy) else (NONE_EQ, thy) end end local fun no_no_code ctxt (rty, qty) = if same_type_constrs (rty, qty) then forall (no_no_code ctxt) (Targs rty ~~ Targs qty) else if Term.is_Type qty then if Lifting_Info.is_no_code_type ctxt (Tname qty) then false else let val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty) val (rty's, rtyqs) = (Targs rty', Targs rtyq) in forall (no_no_code ctxt) (rty's ~~ rtyqs) end else true fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = let fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term in Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty] end exception DECODE fun decode_code_eq thm = if Thm.nprems_of thm > 0 then raise DECODE else let val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type in (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) end structure Data = Generic_Data ( type T = code_eq option val empty = NONE fun merge _ = NONE ); fun register_encoded_code_eq thm thy = let val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy in Context.theory_map (Data.put (SOME code_eq)) thy end handle DECODE => thy val register_code_eq_attribute = Thm.declaration_attribute (fn thm => Context.mapping (register_encoded_code_eq thm) I) - val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute) + val register_code_eq_attrib = Attrib.internal \<^here> (K register_code_eq_attribute) in fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy = let val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty) in if no_no_code lthy (rty, qty) then let val lthy' = lthy |> (#2 oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy')) val code_eq = if is_some opt_code_eq then the opt_code_eq else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know which code equation is going to be used. This is going to be resolved at the point when an interpretation of the locale is executed. *) val lthy'' = lthy' - |> Local_Theory.declaration {syntax = false, pervasive = true} (K (Data.put NONE)) + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (K (Data.put NONE)) in (code_eq, lthy'') end else (NONE_EQ, lthy) end end (* Defines an operation on an abstract type in terms of a corresponding operation on a representation type. var - a binding and a mixfix of the new constant being defined qty - an abstract type of the new constant rhs - a term representing the new constant on the raw level rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'), i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs" par_thms - a parametricity theorem for rhs *) fun add_lift_def (config: config) (binding, mx) qty rhs rsp_thm par_thms lthy0 = let val rty = fastype_of rhs val quot_thm = Lifting_Term.prove_quot_thm lthy0 (rty, qty) val absrep_trm = quot_thm_abs quot_thm val rty_forced = (domain_type o fastype_of) absrep_trm val forced_rhs = force_rty_type lthy0 rty_forced rhs val lhs = Free (Binding.name_of binding, qty) val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs) val (_, prop') = Local_Defs.cert_def lthy0 (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val var = ((#notes config = false ? Binding.concealed) binding, mx) val def_name = Thm.make_def_binding (#notes config) (#1 var) val ((lift_const, (_ , def_thm)), lthy1) = lthy0 |> Local_Theory.define (var, ((def_name, []), newrhs)) val transfer_rules = generate_transfer_rules lthy1 quot_thm rsp_thm def_thm par_thms val abs_eq_thm = generate_abs_eq lthy1 def_thm rsp_thm quot_thm val opt_rep_eq_thm = generate_rep_eq lthy1 def_thm rsp_thm (rty_forced, qty) fun notes names = let val lhs_name = Binding.reset_pos (#1 var) val rsp_thmN = Binding.qualify_name true lhs_name "rsp" val abs_eq_thmN = Binding.qualify_name true lhs_name "abs_eq" val rep_eq_thmN = Binding.qualify_name true lhs_name "rep_eq" val transfer_ruleN = Binding.qualify_name true lhs_name "transfer" val notes = [(rsp_thmN, [], [rsp_thm]), (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules), (abs_eq_thmN, [], [abs_eq_thm])] @ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => []) in if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes else map_filter (fn (_, attrs, thms) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) notes end val (code_eq, lthy2) = lthy1 |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm opt_rep_eq_thm code_eq transfer_rules in lthy2 |> (snd o Local_Theory.begin_nested) |> Local_Theory.notes (notes (#notes config)) |> snd |> `(fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def) ||> Local_Theory.end_nested end (* This is not very cheap way of getting the rules but we have only few active liftings in the current setting *) fun get_cr_pcr_eqs ctxt = let fun collect (data : Lifting_Info.quotient) l = if is_some (#pcr_info data) then ((Thm.symmetric o safe_mk_meta_eq o Thm.transfer' ctxt o #pcr_cr_eq o the o #pcr_info) data :: l) else l val table = Lifting_Info.get_quotients ctxt in Symtab.fold (fn (_, data) => fn l => collect data l) table [] end fun prepare_lift_def add_lift_def var qty rhs par_thms ctxt = let val rsp_rel = Lifting_Term.equiv_relation ctxt (fastype_of rhs, qty) val rty_forced = (domain_type o fastype_of) rsp_rel; val forced_rhs = force_rty_type ctxt rty_forced rhs; val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (Raw_Simplifier.rewrite ctxt false (get_cr_pcr_eqs ctxt))) val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs) |> Thm.cterm_of ctxt |> cr_to_pcr_conv |> `Thm.concl_of |>> Logic.dest_equals |>> snd val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2 val opt_proven_rsp_thm = try_prove_reflexivity ctxt prsp_tm fun after_qed internal_rsp_thm = add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms in case opt_proven_rsp_thm of SOME thm => (NONE, K (after_qed thm)) | NONE => (SOME prsp_tm, after_qed) end fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy = let val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy in case goal of SOME goal => let val rsp_thm = Goal.prove_sorry lthy [] [] goal (tac o #context) |> Thm.close_derivation \<^here> in after_qed rsp_thm lthy end | NONE => after_qed Drule.dummy_thm lthy end val lift_def = gen_lift_def o add_lift_def end (* structure *) diff --git a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML --- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML +++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML @@ -1,833 +1,833 @@ (* Title: HOL/Tools/Lifting/lifting_def_code_dt.ML Author: Ondrej Kuncar Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype; lift_definition command *) signature LIFTING_DEF_CODE_DT = sig type rep_isom_data val isom_of_rep_isom_data: rep_isom_data -> term val transfer_of_rep_isom_data: rep_isom_data -> thm val bundle_name_of_rep_isom_data: rep_isom_data -> string val pointer_of_rep_isom_data: rep_isom_data -> string type code_dt val rty_of_code_dt: code_dt -> typ val qty_of_code_dt: code_dt -> typ val wit_of_code_dt: code_dt -> term val wit_thm_of_code_dt: code_dt -> thm val rep_isom_data_of_code_dt: code_dt -> rep_isom_data option val morph_code_dt: morphism -> code_dt -> code_dt val mk_witness_of_code_dt: typ -> code_dt -> term val mk_rep_isom_of_code_dt: typ -> code_dt -> term option val code_dt_of: Proof.context -> typ * typ -> code_dt option val code_dt_of_global: theory -> typ * typ -> code_dt option val all_code_dt_of: Proof.context -> code_dt list val all_code_dt_of_global: theory -> code_dt list type config_code_dt = { code_dt: bool, lift_config: Lifting_Def.config } val default_config_code_dt: config_code_dt val add_lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_cmd: string list * (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list -> local_theory -> Proof.state end structure Lifting_Def_Code_Dt: LIFTING_DEF_CODE_DT = struct open Ctr_Sugar_Util BNF_Util BNF_FP_Util BNF_FP_Def_Sugar Lifting_Def Lifting_Util infix 0 MRSL (** data structures **) (* all type variables in qty are in rty *) datatype rep_isom_data = REP_ISOM of { isom: term, transfer: thm, bundle_name: string, pointer: string } fun isom_of_rep_isom_data (REP_ISOM rep_isom) = #isom rep_isom; fun transfer_of_rep_isom_data (REP_ISOM rep_isom) = #transfer rep_isom; fun bundle_name_of_rep_isom_data (REP_ISOM rep_isom) = #bundle_name rep_isom; fun pointer_of_rep_isom_data (REP_ISOM rep_isom) = #pointer rep_isom; datatype code_dt = CODE_DT of { rty: typ, qty: typ, wit: term, wit_thm: thm, rep_isom_data: rep_isom_data option }; fun rty_of_code_dt (CODE_DT code_dt) = #rty code_dt; fun qty_of_code_dt (CODE_DT code_dt) = #qty code_dt; fun wit_of_code_dt (CODE_DT code_dt) = #wit code_dt; fun wit_thm_of_code_dt (CODE_DT code_dt) = #wit_thm code_dt; fun rep_isom_data_of_code_dt (CODE_DT code_dt) = #rep_isom_data code_dt; fun ty_alpha_equiv (T, U) = Type.raw_instance (T, U) andalso Type.raw_instance (U, T); fun code_dt_eq c = (ty_alpha_equiv o apply2 rty_of_code_dt) c andalso (ty_alpha_equiv o apply2 qty_of_code_dt) c; fun term_of_code_dt code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> Net.encode_type |> single; (* modulo renaming, typ must contain TVars *) fun is_code_dt_of_type (rty, qty) code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> curry ty_alpha_equiv (HOLogic.mk_prodT (rty, qty)); fun mk_rep_isom_data isom transfer bundle_name pointer = REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer} fun mk_code_dt rty qty wit wit_thm rep_isom_data = CODE_DT { rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data }; fun map_rep_isom_data f1 f2 f3 f4 (REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer }) = REP_ISOM { isom = f1 isom, transfer = f2 transfer, bundle_name = f3 bundle_name, pointer = f4 pointer }; fun map_code_dt f1 f2 f3 f4 f5 f6 f7 f8 (CODE_DT {rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data}) = CODE_DT {rty = f1 rty, qty = f2 qty, wit = f3 wit, wit_thm = f4 wit_thm, rep_isom_data = Option.map (map_rep_isom_data f5 f6 f7 f8) rep_isom_data}; fun update_rep_isom isom transfer binding pointer i = mk_code_dt (rty_of_code_dt i) (qty_of_code_dt i) (wit_of_code_dt i) (wit_thm_of_code_dt i) (SOME (mk_rep_isom_data isom transfer binding pointer)) fun morph_code_dt phi = let val mty = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_code_dt mty mty mterm mthm mterm mthm I I end val transfer_code_dt = morph_code_dt o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = code_dt Item_Net.T val empty = Item_Net.init code_dt_eq term_of_code_dt val merge = Item_Net.merge ); fun code_dt_of_generic context (rty, qty) = let val typ = HOLogic.mk_prodT (rty, qty) val prefiltred = Item_Net.retrieve_matching (Data.get context) (Net.encode_type typ) in prefiltred |> filter (is_code_dt_of_type (rty, qty)) |> map (transfer_code_dt (Context.theory_of context)) |> find_first (fn _ => true) end; fun code_dt_of ctxt (rty, qty) = let val sch_rty = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty val sch_qty = Logic.type_map (singleton (Variable.polymorphic ctxt)) qty in code_dt_of_generic (Context.Proof ctxt) (sch_rty, sch_qty) end; fun code_dt_of_global thy (rty, qty) = let val sch_rty = Logic.varifyT_global rty val sch_qty = Logic.varifyT_global qty in code_dt_of_generic (Context.Theory thy) (sch_rty, sch_qty) end; fun all_code_dt_of_generic context = Item_Net.content (Data.get context) |> map (transfer_code_dt (Context.theory_of context)); val all_code_dt_of = all_code_dt_of_generic o Context.Proof; val all_code_dt_of_global = all_code_dt_of_generic o Context.Theory; fun update_code_dt code_dt = (snd o Local_Theory.begin_nested) - #> Local_Theory.declaration {syntax = false, pervasive = true} + #> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Item_Net.update (morph_code_dt phi code_dt))) #> Local_Theory.end_nested fun mk_match_of_code_dt qty code_dt = Vartab.empty |> Type.raw_match (qty_of_code_dt code_dt, qty) |> Vartab.dest |> map (fn (x, (S, T)) => (TVar (x, S), T)); fun mk_witness_of_code_dt qty code_dt = Term.subst_atomic_types (mk_match_of_code_dt qty code_dt) (wit_of_code_dt code_dt) fun mk_rep_isom_of_code_dt qty code_dt = Option.map (isom_of_rep_isom_data #> Term.subst_atomic_types (mk_match_of_code_dt qty code_dt)) (rep_isom_data_of_code_dt code_dt) (** unique name for a type **) fun var_name name sort = if sort = \<^sort>\{type}\ orelse sort = [] then ["x" ^ name] else "x" ^ name :: "x_" :: sort @ ["x_"]; fun concat_Tnames (Type (name, ts)) = name :: maps concat_Tnames ts | concat_Tnames (TFree (name, sort)) = var_name name sort | concat_Tnames (TVar ((name, _), sort)) = var_name name sort; fun unique_Tname (rty, qty) = let val Tnames = map Long_Name.base_name (concat_Tnames rty @ ["x_x"] @ concat_Tnames qty); in fold (Binding.qualify false) (tl Tnames) (Binding.name (hd Tnames)) end; (** witnesses **) fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_witness quot_thm = let val wit_thm = quot_thm RS @{thm type_definition_Quotient_not_empty_witness} val wit = quot_thm_rep quot_thm $ mk_undefined (quot_thm_rty_qty quot_thm |> snd) in (wit, wit_thm) end (** config **) type config_code_dt = { code_dt: bool, lift_config: config } val default_config_code_dt = { code_dt = false, lift_config = default_config } (** Main code **) val ld_no_notes = { notes = false } fun comp_lift_error _ _ = error "Composition of abstract types has not been implemented yet." fun lift qty (quot_thm, (lthy, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm; in if is_none (code_dt_of lthy (rty, qty)) then let val (wit, wit_thm) = (mk_witness quot_thm handle THM _ => error ("code_dt: " ^ quote (Tname qty) ^ " was not defined as a subtype.")) val code_dt = mk_code_dt rty qty wit wit_thm NONE in (quot_thm, (update_code_dt code_dt lthy, rel_eq_onps)) end else (quot_thm, (lthy, rel_eq_onps)) end; fun case_tac rule = Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} => HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME (snd (hd params))] rule))); fun bundle_name_of_bundle_binding binding phi context = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi binding); fun prove_schematic_quot_thm actions ctxt = Lifting_Term.prove_schematic_quot_thm actions (Lifting_Info.get_quotients ctxt) ctxt fun prove_code_dt (rty, qty) lthy = let val (fold_quot_thm: (local_theory * thm list) Lifting_Term.fold_quot_thm) = { constr = constr, lift = lift, comp_lift = comp_lift_error }; in prove_schematic_quot_thm fold_quot_thm lthy (rty, qty) (lthy, []) |> snd end and add_lift_def_code_dt config var qty rhs rsp_thm par_thms lthy = let fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun ret_rel_conv conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => binop_conv2 Conv.all_conv conv ctm | _ => conv ctm fun R_conv rel_eq_onps ctxt = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} ctxt then_conv Conv.bottom_rewrs_conv rel_eq_onps ctxt val (ret_lift_def, lthy1) = add_lift_def (#lift_config config) var qty rhs rsp_thm par_thms lthy in if (not (#code_dt config) orelse (code_eq_of_lift_def ret_lift_def <> NONE_EQ) andalso (code_eq_of_lift_def ret_lift_def <> UNKNOWN_EQ)) (* Let us try even in case of UNKNOWN_EQ. If this leads to problems, the user can always say that they do not want this workaround. *) then (ret_lift_def, lthy1) else let val lift_def = inst_of_lift_def lthy1 qty ret_lift_def val rty = rty_of_lift_def lift_def val rty_ret = body_type rty val qty_ret = body_type qty val (lthy2, rel_eq_onps) = prove_code_dt (rty_ret, qty_ret) lthy1 val code_dt = code_dt_of lthy2 (rty_ret, qty_ret) in if is_none code_dt orelse is_none (rep_isom_data_of_code_dt (the code_dt)) then (ret_lift_def, lthy2) else let val code_dt = the code_dt val rhs = dest_comb (rhs_of_lift_def lift_def) |> snd val rep_isom_data = code_dt |> rep_isom_data_of_code_dt |> the val pointer = pointer_of_rep_isom_data rep_isom_data val quot_active = Lifting_Info.lookup_restore_data lthy2 pointer |> the |> #quotient |> #quot_thm |> Lifting_Info.lookup_quot_thm_quotients lthy2 |> is_some val qty_code_dt_bundle_name = bundle_name_of_rep_isom_data rep_isom_data val rep_isom = mk_rep_isom_of_code_dt qty_ret code_dt |> the val lthy3 = if quot_active then lthy2 else Bundle.includes [qty_code_dt_bundle_name] lthy2 fun qty_isom_of_rep_isom rep = rep |> dest_Const |> snd |> domain_type val qty_isom = qty_isom_of_rep_isom rep_isom val f'_var = (Binding.suffix_name "_aux" (fst var), NoSyn); val f'_qty = strip_type qty |> fst |> rpair qty_isom |> op ---> val f'_rsp_rel = Lifting_Term.equiv_relation lthy3 (rty, f'_qty); val rsp = rsp_thm_of_lift_def lift_def val rel_eq_onps_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (ret_rel_conv (R_conv rel_eq_onps lthy3))) val rsp_norm = Conv.fconv_rule rel_eq_onps_conv rsp val f'_rsp_goal = HOLogic.mk_Trueprop (f'_rsp_rel $ rhs $ rhs); val f'_rsp = Goal.prove_sorry lthy3 [] [] f'_rsp_goal (fn {context = ctxt, prems = _} => HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac ctxt rsp_norm)) |> Thm.close_derivation \<^here> val (f'_lift_def, lthy4) = add_lift_def ld_no_notes f'_var f'_qty rhs f'_rsp [] lthy3 val f'_lift_def = inst_of_lift_def lthy4 f'_qty f'_lift_def val f'_lift_const = mk_lift_const_of_lift_def f'_qty f'_lift_def val (args, args_ctxt) = mk_Frees "x" (binder_types qty) lthy4 val f_alt_def_goal_lhs = list_comb (lift_const_of_lift_def lift_def, args); val f_alt_def_goal_rhs = rep_isom $ list_comb (f'_lift_const, args); val f_alt_def_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (f_alt_def_goal_lhs, f_alt_def_goal_rhs)); fun f_alt_def_tac ctxt i = EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [id_apply]), rtac ctxt refl] i; val rep_isom_transfer = transfer_of_rep_isom_data rep_isom_data val (_, transfer_ctxt) = args_ctxt |> Proof_Context.note_thms "" (Binding.empty_atts, [([rep_isom_transfer], [Transfer.transfer_add])]) val f_alt_def = Goal.prove_sorry transfer_ctxt [] [] f_alt_def_goal (fn {context = goal_ctxt, ...} => HEADGOAL (f_alt_def_tac goal_ctxt)) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy4) val lthy5 = lthy4 |> Local_Theory.note ((Binding.empty, @{attributes [code]}), [f_alt_def]) |> snd (* if processing a mutual datatype (there is a cycle!) the corresponding quotient will be needed later and will be forgotten later *) |> (if quot_active then I else Lifting_Setup.lifting_forget pointer) in (ret_lift_def, lthy5) end end end and mk_rep_isom qty_isom_bundle (rty, qty, qty_isom) lthy0 = let (* logical definition of qty qty_isom isomorphism *) val uTname = unique_Tname (rty, qty) fun eq_onp_to_top_tac ctxt = SELECT_GOAL (Local_Defs.unfold0_tac ctxt (@{thm eq_onp_top_eq_eq[symmetric]} :: Lifting_Info.get_relator_eq_onp_rules ctxt)) fun lift_isom_tac ctxt = HEADGOAL (eq_onp_to_top_tac ctxt THEN' (rtac ctxt @{thm id_transfer})); val (rep_isom_lift_def, lthy1) = lthy0 |> (snd o Local_Theory.begin_nested) |> lift_def ld_no_notes (Binding.qualify_name true uTname "Rep_isom", NoSyn) (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] |>> inst_of_lift_def lthy0 (qty_isom --> qty); val (abs_isom, lthy2) = lthy1 |> lift_def ld_no_notes (Binding.qualify_name true uTname "Abs_isom", NoSyn) (qty --> qty_isom) (HOLogic.id_const rty) lift_isom_tac [] |>> mk_lift_const_of_lift_def (qty --> qty_isom); val rep_isom = lift_const_of_lift_def rep_isom_lift_def val pointer = Lifting_Setup.pointer_of_bundle_binding lthy2 qty_isom_bundle fun code_dt phi context = code_dt_of lthy2 (rty, qty) |> the |> update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd) (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer; val lthy3 = lthy2 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context) |> Local_Theory.end_nested (* in order to make the qty qty_isom isomorphism executable we have to define discriminators and selectors for qty_isom *) val (rty_name, typs) = dest_Type rty val (_, qty_typs) = dest_Type qty val fp = BNF_FP_Def_Sugar.fp_sugar_of lthy3 rty_name val fp = if is_some fp then the fp else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val ctr_sugar = fp |> #fp_ctr_sugar |> #ctr_sugar val ctrs = map (Ctr_Sugar.mk_ctr typs) (#ctrs ctr_sugar); val qty_ctrs = map (Ctr_Sugar.mk_ctr qty_typs) (#ctrs ctr_sugar); val ctr_Tss = map (dest_Const #> snd #> binder_types) ctrs; val qty_ctr_Tss = map (dest_Const #> snd #> binder_types) qty_ctrs; val n = length ctrs; val ks = 1 upto n; val (xss, _) = mk_Freess "x" ctr_Tss lthy3; fun sel_retT (rty' as Type (s, rtys'), qty' as Type (s', qtys')) = if (rty', qty') = (rty, qty) then qty_isom else (if s = s' then Type (s, map sel_retT (rtys' ~~ qtys')) else qty') | sel_retT (_, qty') = qty'; val sel_retTs = map2 (map2 (sel_retT oo pair)) ctr_Tss qty_ctr_Tss fun lazy_prove_code_dt (rty, qty) lthy = if is_none (code_dt_of lthy (rty, qty)) then prove_code_dt (rty, qty) lthy |> fst else lthy; val lthy4 = fold2 (fold2 (lazy_prove_code_dt oo pair)) ctr_Tss sel_retTs lthy3 val sel_argss = @{map 4} (fn k => fn xs => @{map 2} (fn x => fn qty_ret => (k, qty_ret, (xs, x)))) ks xss xss sel_retTs; fun mk_sel_casex (_, _, (_, x)) = Ctr_Sugar.mk_case typs (x |> dest_Free |> snd) (#casex ctr_sugar); val dis_casex = Ctr_Sugar.mk_case typs HOLogic.boolT (#casex ctr_sugar); fun mk_sel_case_args lthy ctr_Tss ks (k, qty_ret, (xs, x)) = let val T = x |> dest_Free |> snd; fun gen_undef_wit Ts wits = case code_dt_of lthy (T, qty_ret) of SOME code_dt => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_witness_of_code_dt qty_ret code_dt), wit_thm_of_code_dt code_dt :: wits) | NONE => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T), wits) in @{fold_map 2} (fn Ts => fn k' => fn wits => (if k = k' then (fold_rev Term.lambda xs x, wits) else gen_undef_wit Ts wits)) ctr_Tss ks [] end; fun mk_sel_rhs arg = let val (sel_rhs, wits) = mk_sel_case_args lthy4 ctr_Tss ks arg in (arg |> #2, wits, list_comb (mk_sel_casex arg, sel_rhs)) end; fun mk_dis_case_args args k = map (fn (k', arg) => (if k = k' then fold_rev Term.lambda arg \<^Const>\True\ else fold_rev Term.lambda arg \<^Const>\False\)) args; val sel_rhs = map (map mk_sel_rhs) sel_argss val dis_rhs = map (fn k => list_comb (dis_casex, mk_dis_case_args (ks ~~ xss) k)) ks val dis_qty = qty_isom --> HOLogic.boolT; val dis_names = map (fn k => Binding.qualify_name true uTname ("dis" ^ string_of_int k)) ks; val (diss, lthy5) = @{fold_map 2} (fn b => fn rhs => fn lthy => lift_def ld_no_notes (b, NoSyn) dis_qty rhs (K all_tac) [] lthy |>> mk_lift_const_of_lift_def dis_qty) dis_names dis_rhs lthy4 val unfold_lift_sel_rsp = @{lemma "(\x. P1 x \ P2 (f x)) \ (rel_fun (eq_onp P1) (eq_onp P2)) f f" by (simp add: eq_onp_same_args rel_fun_eq_onp_rel)} fun lift_sel_tac exhaust_rule dt_rules wits ctxt i = (Method.insert_tac ctxt wits THEN' eq_onp_to_top_tac ctxt THEN' (* normalize *) rtac ctxt unfold_lift_sel_rsp THEN' case_tac exhaust_rule ctxt THEN_ALL_NEW ( EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *) Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt])) i val pred_simps = Transfer.lookup_pred_data lthy5 (Tname rty) |> the |> Transfer.pred_simps val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps) val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualify_name true uTname ("sel" ^ string_of_int k ^ string_of_int k')) (1 upto length xs)) (ks ~~ ctr_Tss); val (selss, lthy6) = @{fold_map 2} (@{fold_map 2} (fn b => fn (qty_ret, wits, rhs) => fn lthy => lift_def_code_dt { code_dt = true, lift_config = ld_no_notes } (b, NoSyn) (qty_isom --> qty_ret) rhs (HEADGOAL o sel_tac wits) [] lthy |>> mk_lift_const_of_lift_def (qty_isom --> qty_ret))) sel_names sel_rhs lthy5 (* now we can execute the qty qty_isom isomorphism *) fun mk_type_definition newT oldT RepC AbsC A = let val typedefC = Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in typedefC $ RepC $ AbsC $ A end; val typedef_goal = mk_type_definition qty_isom qty rep_isom abs_isom (HOLogic.mk_UNIV qty) |> HOLogic.mk_Trueprop; fun typ_isom_tac ctxt i = EVERY' [ SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms type_definition_def}), DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms eq_onp_top_eq_eq}) (* normalize *), Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}), rtac ctxt TrueI] i; val (_, transfer_ctxt) = Proof_Context.note_thms "" (Binding.empty_atts, [(@{thms right_total_UNIV_transfer}, [Transfer.transfer_add]), (@{thms Domain_eq_top}, [Transfer.transfer_domain_add])]) lthy6; val quot_thm_isom = Goal.prove_sorry transfer_ctxt [] [] typedef_goal (fn {context = goal_ctxt, ...} => typ_isom_tac goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy6) |> (fn thm => @{thm UNIV_typedef_to_Quotient} OF [thm, @{thm reflexive}]) val qty_isom_name = Tname qty_isom; val quot_isom_rep = let val (quotients : Lifting_Term.quotients) = Symtab.insert (Lifting_Info.quotient_eq) (qty_isom_name, {quot_thm = quot_thm_isom, pcr_info = NONE}) Symtab.empty val id_actions = { constr = K I, lift = K I, comp_lift = K I } in fn ctxt => fn (rty, qty) => Lifting_Term.prove_schematic_quot_thm id_actions quotients ctxt (rty, qty) () |> fst |> Lifting_Term.force_qty_type ctxt qty |> quot_thm_rep end; val (x, x_ctxt) = yield_singleton (mk_Frees "x") qty_isom lthy6; fun mk_ctr ctr ctr_Ts sels = let val sel_ret_Ts = map (dest_Const #> snd #> body_type) sels; fun rep_isom lthy t (rty, qty) = let val rep = quot_isom_rep lthy (rty, qty) in if is_Const rep andalso (rep |> dest_Const |> fst) = \<^const_name>\id\ then t else rep $ t end; in @{fold 3} (fn sel => fn ctr_T => fn sel_ret_T => fn ctr => ctr $ rep_isom x_ctxt (sel $ x) (ctr_T, sel_ret_T)) sels ctr_Ts sel_ret_Ts ctr end; (* stolen from Metis *) exception BREAK_LIST fun break_list (x :: xs) = (x, xs) | break_list _ = raise BREAK_LIST val (ctr, ctrs) = qty_ctrs |> rev |> break_list; val (ctr_Ts, ctr_Tss) = qty_ctr_Tss |> rev |> break_list; val (sel, rselss) = selss |> rev |> break_list; val rdiss = rev diss |> tl; val first_ctr = mk_ctr ctr ctr_Ts sel; fun mk_If_ctr dis ctr ctr_Ts sel elsex = mk_If (dis$x) (mk_ctr ctr ctr_Ts sel) elsex; val rhs = @{fold 4} mk_If_ctr rdiss ctrs ctr_Tss rselss first_ctr; val rep_isom_code_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (rep_isom$x, rhs)); local val rep_isom_code_tac_rules = map safe_mk_meta_eq @{thms refl id_apply if_splits simp_thms} in fun rep_isom_code_tac (ctr_sugar:Ctr_Sugar.ctr_sugar) ctxt i = let val exhaust = ctr_sugar |> #exhaust val cases = ctr_sugar |> #case_thms val map_ids = fp |> #fp_nesting_bnfs |> map BNF_Def.map_id0_of_bnf val simp_rules = map safe_mk_meta_eq (cases @ map_ids) @ rep_isom_code_tac_rules in EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o (Transfer.transfer_tac true ctxt), case_tac exhaust ctxt THEN_ALL_NEW EVERY' [hyp_subst_tac ctxt, Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac ctxt TrueI ]] i end end (* stolen from bnf_fp_n2m.ML *) fun force_typ ctxt T = Term.map_types Type_Infer.paramify_vars #> Type.constraint T #> singleton (Type_Infer_Context.infer_types ctxt); (* The following tests that types in rty have corresponding arities imposed by constraints of the datatype fp. Otherwise rep_isom_code_tac could fail (especially transfer in it) is such a way that it is not easy to infer the problem with sorts. *) val _ = yield_singleton (mk_Frees "x") (#T fp) x_ctxt |> fst |> force_typ x_ctxt qty val rep_isom_code = Goal.prove_sorry x_ctxt [] [] rep_isom_code_goal (fn {context = goal_ctxt, ...} => rep_isom_code_tac ctr_sugar goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton(Variable.export x_ctxt lthy6) in lthy6 |> snd o Local_Theory.note ((Binding.empty, @{attributes [code]}), [rep_isom_code]) |> Lifting_Setup.lifting_forget pointer |> pair (selss, diss, rep_isom_code) end and constr qty (quot_thm, (lthy0, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy0 qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm val rty_name = Tname rty; val pred_data = Transfer.lookup_pred_data lthy0 rty_name val pred_data = if is_some pred_data then the pred_data else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val rel_eq_onp = safe_mk_meta_eq (Transfer.rel_eq_onp pred_data); val rel_eq_onps = insert Thm.eq_thm rel_eq_onp rel_eq_onps fun R_conv ctxt = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} ctxt then_conv Conv.rewr_conv rel_eq_onp val quot_thm = Conv.fconv_rule (HOLogic.Trueprop_conv (Quotient_R_conv (R_conv lthy0))) quot_thm; in if is_none (code_dt_of lthy0 (rty, qty)) then let val non_empty_pred = quot_thm RS @{thm type_definition_Quotient_not_empty} val pred = quot_thm_rel quot_thm |> dest_comb |> snd; val (pred, lthy1) = lthy0 |> (snd o Local_Theory.begin_nested) |> yield_singleton (Variable.import_terms true) pred; val TFrees = Term.add_tfreesT qty [] fun non_empty_typedef_tac non_empty_pred ctxt i = (Method.insert_tac ctxt [non_empty_pred] THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt [mem_Collect_eq]) THEN' assume_tac ctxt) i val uTname = unique_Tname (rty, qty) val Tdef_set = HOLogic.mk_Collect ("x", rty, pred $ Free("x", rty)); val ((_, tcode_dt), lthy2) = lthy1 |> conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn) Tdef_set NONE (fn lthy => HEADGOAL (non_empty_typedef_tac non_empty_pred lthy))); val type_definition_thm = tcode_dt |> snd |> #type_definition; val qty_isom = tcode_dt |> fst |> #abs_type; val (binding, lthy3) = lthy2 |> conceal_naming_result (Lifting_Setup.setup_by_typedef_thm {notes = false} type_definition_thm) ||> Local_Theory.end_nested val (wit, wit_thm) = mk_witness quot_thm; val code_dt = mk_code_dt rty qty wit wit_thm NONE; val lthy4 = lthy3 |> update_code_dt code_dt |> mk_rep_isom binding (rty, qty, qty_isom) |> snd in (quot_thm, (lthy4, rel_eq_onps)) end else (quot_thm, (lthy0, rel_eq_onps)) end and lift_def_code_dt config = gen_lift_def (add_lift_def_code_dt config) (** from parsed parameters to the config record **) fun map_config_code_dt f1 f2 ({code_dt = code_dt, lift_config = lift_config}: config_code_dt) = {code_dt = f1 code_dt, lift_config = f2 lift_config} fun update_config_code_dt nval = map_config_code_dt (K nval) I val config_flags = [("code_dt", update_config_code_dt true)] fun evaluate_params params = let fun eval_param param config = case AList.lookup (op =) config_flags param of SOME update => update config | NONE => error ("Unknown parameter: " ^ (quote param)) in fold eval_param params default_config_code_dt end (** lift_definition command. It opens a proof of a corresponding respectfulness theorem in a user-friendly, readable form. Then add_lift_def_code_dt is called internally. **) local val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm \<^context>) [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, @{thm pcr_Domainp}] in fun mk_readable_rsp_thm_eq tm ctxt = let val ctm = Thm.cterm_of ctxt tm fun assms_rewr_conv tactic rule ct = let fun prove_extra_assms thm = let val assms = cprems_of thm fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm)) in map_interrupt prove assms end fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm) fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm)) fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm)) val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule; val lhs = lhs_of rule1; val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1; val rule3 = Thm.instantiate (Thm.match (lhs, ct)) rule2 handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]); val proved_assms = prove_extra_assms rule3 in case proved_assms of SOME proved_assms => let val rule3 = proved_assms MRSL rule3 val rule4 = if lhs_of rule3 aconvc ct then rule3 else let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3) in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end | NONE => Conv.no_conv ct end fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules) fun simp_arrows_conv ctm = let val unfold_conv = Conv.rewrs_conv [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_onp_rel[THEN eq_reflection]}, @{thm rel_fun_eq[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, @{thm rel_fun_def[THEN eq_reflection]}] fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw ctxt) val intro_top_rule = @{thm eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} val kill_tops = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[THEN eq_reflection]} ctxt val eq_onp_assms_tac = (CONVERSION kill_tops THEN' TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_onp_assms_tac_rules) THEN_ALL_NEW (DETERM o Transfer.eq_tac ctxt)) 1 val relator_eq_onp_conv = Conv.bottom_conv (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac (intro_top_rule :: Lifting_Info.get_relator_eq_onp_rules ctxt)))) ctxt then_conv kill_tops val relator_eq_conv = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq ctxt)))) ctxt in case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm end val unfold_ret_val_invs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt val unfold_inv_conv = Conv.top_sweep_rewrs_conv @{thms eq_onp_def[THEN eq_reflection]} ctxt val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv simp_arrows_conv) val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt val beta_conv = Thm.beta_conversion true val eq_thm = (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs then_conv unfold_inv_conv) ctm in Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) end end fun rename_to_tnames ctxt term = let fun all_typs (Const (\<^const_name>\Pure.all\, _) $ Abs (_, T, t)) = T :: all_typs t | all_typs _ = [] fun rename (Const (\<^const_name>\Pure.all\, T1) $ Abs (_, T2, t)) (new_name :: names) = (Const (\<^const_name>\Pure.all\, T1) $ Abs (new_name, T2, rename t names)) | rename t _ = t val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t) in rename term new_names end fun quot_thm_err ctxt (rty, qty) pretty_msg = let val error_msg = cat_lines ["Lifting failed for the following types:", Pretty.string_of (Pretty.block [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]), Pretty.string_of (Pretty.block [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))] in error error_msg end fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) = let val (_, ctxt') = Proof_Context.read_var raw_var ctxt val rhs = Syntax.read_term ctxt' rhs_raw val error_msg = cat_lines ["Lifting failed for the following term:", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, Pretty.str "The type of the term cannot be instantiated to", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt rty_forced), Pretty.str "."]))] in error error_msg end fun lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy0 = let val config = evaluate_params params val ((binding, SOME qty, mx), lthy1) = Proof_Context.read_var raw_var lthy0 val var = (binding, mx) val rhs = Syntax.read_term lthy1 rhs_raw val par_thms = Attrib.eval_thms lthy1 par_xthms val (goal, after_qed) = lthy1 |> prepare_lift_def (add_lift_def_code_dt config) var qty rhs par_thms val (goal, after_qed) = case goal of NONE => (goal, K (after_qed Drule.dummy_thm)) | SOME prsp_tm => let val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy1 val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) val readable_rsp_tm_tnames = rename_to_tnames lthy1 readable_rsp_tm fun after_qed' [[thm]] lthy = let val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm (fn {context = goal_ctxt, ...} => rtac goal_ctxt readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac goal_ctxt [thm] 1) in after_qed internal_rsp_thm lthy end in (SOME readable_rsp_tm_tnames, after_qed') end fun after_qed_with_err_handling thmss ctxt = (after_qed thmss ctxt handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy1 (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy1 (rty_schematic, rty_forced) (raw_var, rhs_raw); in lthy1 |> Proof.theorem NONE (snd oo after_qed_with_err_handling) [map (rpair []) (the_list goal)] end fun lift_def_cmd_with_err_handling (params, (raw_var, rhs_raw, par_xthms)) lthy = (lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw); val parse_param = Parse.name val parse_params = Scan.optional (Args.parens (Parse.list parse_param)) []; (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\lift_definition\ "definition for constants over the quotient type" (parse_params -- (((Parse.binding -- (\<^keyword>\::\ |-- (Parse.typ >> SOME) -- Parse.opt_mixfix') >> Scan.triple2) -- (\<^keyword>\is\ |-- Parse.term) -- Scan.optional (\<^keyword>\parametric\ |-- Parse.!!! Parse.thms1) []) >> Scan.triple1) >> lift_def_cmd_with_err_handling); end diff --git a/src/HOL/Tools/Lifting/lifting_setup.ML b/src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML +++ b/src/HOL/Tools/Lifting/lifting_setup.ML @@ -1,1046 +1,1046 @@ (* Title: HOL/Tools/Lifting/lifting_setup.ML Author: Ondrej Kuncar Setting up the lifting infrastructure. *) signature LIFTING_SETUP = sig exception SETUP_LIFTING_INFR of string type config = { notes: bool }; val default_config: config; val setup_by_quotient: config -> thm -> thm option -> thm option -> local_theory -> binding * local_theory val setup_by_typedef_thm: config -> thm -> local_theory -> binding * local_theory val lifting_restore: Lifting_Info.quotient -> Context.generic -> Context.generic val lifting_forget: string -> local_theory -> local_theory val update_transfer_rules: string -> local_theory -> local_theory val pointer_of_bundle_binding: Proof.context -> binding -> string end structure Lifting_Setup: LIFTING_SETUP = struct open Lifting_Util infix 0 MRSL exception SETUP_LIFTING_INFR of string (* Config *) type config = { notes: bool }; val default_config = { notes = true }; fun define_crel (config: config) rep_fun lthy = let val (qty, rty) = (dest_funT o fastype_of) rep_fun val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0) val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph)) val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val crel_name = Binding.prefix_name "cr_" qty_name val (fixed_def_term, lthy1) = lthy |> yield_singleton (Variable.importT_terms) def_term val ((_, (_ , def_thm)), lthy2) = if #notes config then Local_Theory.define ((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy1 else Local_Theory.define ((Binding.concealed crel_name, NoSyn), (Binding.empty_atts, fixed_def_term)) lthy1 in (def_thm, lthy2) end fun print_define_pcrel_warning msg = let val warning_msg = cat_lines ["Generation of a parametrized correspondence relation failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in warning warning_msg end fun define_pcrel (config: config) crel lthy0 = let val (fixed_crel, lthy1) = yield_singleton Variable.importT_terms crel lthy0 val [rty', qty] = (binder_types o fastype_of) fixed_crel val (param_rel, args) = Lifting_Term.generate_parametrized_relator lthy1 rty' val rty_raw = (domain_type o range_type o fastype_of) param_rel val tyenv_match = Sign.typ_match (Proof_Context.theory_of lthy1) (rty_raw, rty') Vartab.empty val param_rel_subst = Envir.subst_term (tyenv_match,Vartab.empty) param_rel val args_subst = map (Envir.subst_term (tyenv_match,Vartab.empty)) args val (instT, lthy2) = lthy1 |> Variable.declare_names fixed_crel |> Variable.importT_inst (param_rel_subst :: args_subst) val args_fixed = (map (Term_Subst.instantiate (instT, Vars.empty))) args_subst val param_rel_fixed = Term_Subst.instantiate (instT, Vars.empty) param_rel_subst val rty = (domain_type o fastype_of) param_rel_fixed val relcomp_op = Const (\<^const_name>\relcompp\, (rty --> rty' --> HOLogic.boolT) --> (rty' --> qty --> HOLogic.boolT) --> rty --> qty --> HOLogic.boolT) val qty_name = (fst o dest_Type) qty val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name) val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT]) val lhs = Library.foldl (op $) ((Free (Binding.name_of pcrel_name, relator_type)), args_fixed) val rhs = relcomp_op $ param_rel_fixed $ fixed_crel val definition_term = Logic.mk_equals (lhs, rhs) fun note_def lthy = Specification.definition (SOME (pcrel_name, SOME relator_type, NoSyn)) [] [] (Binding.empty_atts, definition_term) lthy |>> (snd #> snd); fun raw_def lthy = let val ((_, rhs), prove) = Local_Defs.derived_def lthy (K []) {conditional = true} definition_term; val ((_, (_, raw_th)), lthy') = Local_Theory.define ((Binding.concealed pcrel_name, NoSyn), (Binding.empty_atts, rhs)) lthy; val th = prove lthy' raw_th; in (th, lthy') end val (def_thm, lthy3) = if #notes config then note_def lthy2 else raw_def lthy2 in (SOME def_thm, lthy3) end handle Lifting_Term.PARAM_QUOT_THM (_, msg) => (print_define_pcrel_warning msg; (NONE, lthy0)) local val eq_OO_meta = mk_meta_eq @{thm eq_OO} fun print_generate_pcr_cr_eq_error ctxt term = let val goal = Const (\<^const_name>\HOL.eq\, dummyT) $ term $ Const (\<^const_name>\HOL.eq\, dummyT) val error_msg = cat_lines ["Generation of a pcr_cr_eq failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])), "Most probably a relator_eq rule for one of the involved types is missing."] in error error_msg end in fun define_pcr_cr_eq (config: config) lthy pcr_rel_def = let val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type o List.last o binder_types o fastype_of) lhs val args = (snd o strip_comb) lhs fun make_inst var ctxt = let val typ = snd (relation_types (#2 (dest_Var var))) val sort = Type.sort_of_atyp typ val (fresh_var, ctxt') = yield_singleton Variable.invent_types sort ctxt val inst = (#1 (dest_Var var), Thm.cterm_of ctxt' (HOLogic.eq_const (TFree fresh_var))) in (inst, ctxt') end val (args_inst, args_ctxt) = fold_map make_inst args lthy val pcr_cr_eq = pcr_rel_def |> infer_instantiate args_ctxt args_inst |> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.bottom_rewrs_conv (Transfer.get_relator_eq args_ctxt) args_ctxt))) in case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of Const (\<^const_name>\relcompp\, _) $ Const (\<^const_name>\HOL.eq\, _) $ _ => let val thm = pcr_cr_eq |> Conv.fconv_rule (Conv.arg_conv (Conv.rewr_conv eq_OO_meta)) |> HOLogic.mk_obj_eq |> singleton (Variable.export args_ctxt lthy) val lthy' = lthy |> #notes config ? (Local_Theory.note ((Binding.qualify_name true qty_name "pcr_cr_eq", []), [thm]) #> #2) in (thm, lthy') end | Const (\<^const_name>\relcompp\, _) $ t $ _ => print_generate_pcr_cr_eq_error args_ctxt t | _ => error "generate_pcr_cr_eq: implementation error" end end fun define_code_constr quot_thm lthy = let val abs = quot_thm_abs quot_thm in if is_Const abs then let val (fixed_abs, lthy') = yield_singleton Variable.importT_terms abs lthy in Local_Theory.background_theory (Code.declare_datatype_global [dest_Const fixed_abs]) lthy' end else lthy end fun define_abs_type quot_thm = Lifting_Def.can_generate_code_cert quot_thm ? Code.declare_abstype (quot_thm RS @{thm Quotient_abs_rep}); local exception QUOT_ERROR of Pretty.T list in fun quot_thm_sanity_check ctxt quot_thm = let val _ = if (Thm.nprems_of quot_thm > 0) then raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem has extra assumptions:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] else () val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient handle TERM _ => raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem is not of the right form:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt val (rty, qty) = quot_thm_rty_qty quot_thm_fixed val rty_tfreesT = Term.add_tfree_namesT rty [] val qty_tfreesT = Term.add_tfree_namesT qty [] val extra_rty_tfrees = case subtract (op =) qty_tfreesT rty_tfreesT of [] => [] | extras => [Pretty.block ([Pretty.str "Extra variables in the raw type:", Pretty.brk 1] @ ((Pretty.commas o map (Pretty.str o quote)) extras) @ [Pretty.str "."])] val not_type_constr = case qty of Type _ => [] | _ => [Pretty.block [Pretty.str "The quotient type ", Pretty.quote (Syntax.pretty_typ ctxt' qty), Pretty.brk 1, Pretty.str "is not a type constructor."]] val errs = extra_rty_tfrees @ not_type_constr in if null errs then () else raise QUOT_ERROR errs end handle QUOT_ERROR errs => error (cat_lines (["Sanity check of the quotient theorem failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end fun lifting_bundle qty_full_name qinfo lthy = let val thy = Proof_Context.theory_of lthy val binding = Binding.qualify_name true (qty_full_name |> Long_Name.base_name |> Binding.name) "lifting" val morphed_binding = Morphism.binding (Local_Theory.target_morphism lthy) binding val bundle_name = Name_Space.full_name (Name_Space.naming_of (Context.Theory thy)) morphed_binding fun phi_qinfo phi = Lifting_Info.transform_quotient phi qinfo val dummy_thm = Thm.transfer thy Drule.dummy_thm val restore_lifting_att = ([dummy_thm], [map Token.make_string0 ["Lifting.lifting_restore_internal", bundle_name]]) in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi)) |> Bundle.bundle ((binding, [restore_lifting_att])) [] |> pair binding end fun setup_lifting_infr config quot_thm opt_reflp_thm lthy = let val _ = quot_thm_sanity_check lthy quot_thm val (_, qty) = quot_thm_rty_qty quot_thm val (pcrel_def, lthy1) = define_pcrel config (quot_thm_crel quot_thm) lthy (**) val pcrel_def = Option.map (Morphism.thm (Local_Theory.target_morphism lthy1)) pcrel_def (**) val (pcr_cr_eq, lthy2) = case pcrel_def of SOME pcrel_def => apfst SOME (define_pcr_cr_eq config lthy1 pcrel_def) | NONE => (NONE, lthy1) val pcr_info = case pcrel_def of SOME pcrel_def => SOME { pcrel_def = pcrel_def, pcr_cr_eq = the pcr_cr_eq } | NONE => NONE val quotients = { quot_thm = quot_thm, pcr_info = pcr_info } val qty_full_name = (fst o dest_Type) qty fun quot_info phi = Lifting_Info.transform_quotient phi quotients - val reflexivity_rule_attr = Attrib.internal (K Lifting_Info.add_reflexivity_rule_attribute) + val reflexivity_rule_attr = Attrib.internal \<^here> (K Lifting_Info.add_reflexivity_rule_attribute) val lthy3 = case opt_reflp_thm of SOME reflp_thm => lthy2 |> (#2 oo Local_Theory.note) ((Binding.empty, [reflexivity_rule_attr]), [reflp_thm RS @{thm reflp_ge_eq}]) |> define_code_constr quot_thm | NONE => lthy2 |> define_abs_type quot_thm in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) |> lifting_bundle qty_full_name quotients end local fun importT_inst_exclude exclude ts ctxt = let val tvars = rev (subtract op= exclude (fold Term.add_tvars ts [])) val (tfrees, ctxt') = Variable.invent_types (map #2 tvars) ctxt in (TVars.make (tvars ~~ map TFree tfrees), ctxt') end fun import_inst_exclude exclude ts ctxt = let val excludeT = fold (Term.add_tvarsT o snd) exclude [] val (instT, ctxt') = importT_inst_exclude excludeT ts ctxt val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (subtract op= exclude (fold Term.add_vars ts []))) val (xs, ctxt'') = Variable.variant_fixes (map (#1 o #1) vars) ctxt' val inst = Vars.make (vars ~~ map Free (xs ~~ map #2 vars)) in ((instT, inst), ctxt'') end fun import_terms_exclude exclude ts ctxt = let val (inst, ctxt') = import_inst_exclude exclude ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end in fun reduce_goal not_fix goal tac ctxt = let val (fixed_goal, ctxt') = yield_singleton (import_terms_exclude not_fix) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE tac init_goal)) end end local val OO_rules = @{thms left_total_OO left_unique_OO right_total_OO right_unique_OO bi_total_OO bi_unique_OO} in fun parametrize_class_constraint ctxt0 pcr_def constraint = let fun generate_transfer_rule pcr_def constraint goal ctxt = let val (fixed_goal, ctxt') = yield_singleton (Variable.import_terms true) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) val rules = Transfer.get_transfer_raw ctxt' val rules = constraint :: OO_rules @ rules val tac = K (Local_Defs.unfold0_tac ctxt' [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt' rules) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal)) end fun make_goal pcr_def constr = let val pred_name = (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.prop_of) constr val arg = (fst o Logic.dest_equals o Thm.prop_of) pcr_def in HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg) end val check_assms = let val right_names = ["right_total", "right_unique", "left_total", "left_unique", "bi_total", "bi_unique"] fun is_right_name name = member op= right_names (Long_Name.base_name name) fun is_trivial_assm (Const (name, _) $ Var (_, _)) = is_right_name name | is_trivial_assm (Const (name, _) $ Free (_, _)) = is_right_name name | is_trivial_assm _ = false in fn thm => let val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm) val thm_name = (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm val non_trivial_assms = filter_out is_trivial_assm prems in if null non_trivial_assms then () else Pretty.block ([Pretty.str "Non-trivial assumptions in ", Pretty.str thm_name, Pretty.str " transfer rule found:", Pretty.brk 1] @ Pretty.commas (map (Syntax.pretty_term ctxt0) non_trivial_assms)) |> Pretty.string_of |> warning end end val goal = make_goal pcr_def constraint val thm = generate_transfer_rule pcr_def constraint goal ctxt0 val _ = check_assms thm in thm end end local val id_unfold = (Conv.rewr_conv (mk_meta_eq @{thm id_def})) in fun generate_parametric_id lthy rty id_transfer_rule = let (* it doesn't raise an exception because it would have already raised it in define_pcrel *) val (quot_thm, _, ctxt') = Lifting_Term.prove_param_quot_thm lthy rty val parametrized_relator = singleton (Variable.export_terms ctxt' lthy) (quot_thm_crel quot_thm) val id_transfer = @{thm id_transfer} |> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1) |> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold)) val var = hd (Term.add_vars (Thm.prop_of id_transfer) []) val inst = [(#1 var, Thm.cterm_of lthy parametrized_relator)] val id_par_thm = infer_instantiate lthy inst id_transfer in Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm end handle Lifting_Term.MERGE_TRANSFER_REL msg => let val error_msg = cat_lines ["Generation of a parametric transfer rule for the abs. or the rep. function failed.", "A non-parametric version will be used.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in (warning error_msg; id_transfer_rule) end end local fun rewrite_first_Domainp_arg rewr_thm thm = Conv.fconv_rule (Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv rewr_thm))))) thm fun fold_Domainp_pcrel pcrel_def thm = let val ct = thm |> Thm.cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg val pcrel_def = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) pcrel_def val thm' = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def]) in rewrite_first_Domainp_arg (Thm.symmetric pcrel_def) thm' end fun reduce_Domainp ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var val reduced_assm = reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end in fun parametrize_domain dom_thm (pcr_info : Lifting_Info.pcr) ctxt0 = let fun reduce_first_assm ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val reduced_assm = reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end val pcr_cr_met_eq = #pcr_cr_eq pcr_info RS @{thm eq_reflection} val pcr_Domainp_eq = rewrite_first_Domainp_arg (Thm.symmetric pcr_cr_met_eq) dom_thm val pcrel_def = #pcrel_def pcr_info val pcr_Domainp_par_left_total = (dom_thm RS @{thm pcr_Domainp_par_left_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_first_assm ctxt0 (Lifting_Info.get_reflexivity_rules ctxt0) val pcr_Domainp_par = (dom_thm RS @{thm pcr_Domainp_par}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt0 (Transfer.get_relator_domain ctxt0) val pcr_Domainp = (dom_thm RS @{thm pcr_Domainp}) |> fold_Domainp_pcrel pcrel_def val thms = [("domain", [pcr_Domainp], @{attributes [transfer_domain_rule]}), ("domain_par", [pcr_Domainp_par], @{attributes [transfer_domain_rule]}), ("domain_par_left_total", [pcr_Domainp_par_left_total], @{attributes [transfer_domain_rule]}), ("domain_eq", [pcr_Domainp_eq], @{attributes [transfer_domain_rule]})] in thms end fun parametrize_total_domain left_total pcrel_def ctxt = let val thm = (left_total RS @{thm pcr_Domainp_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt) in [("domain", [thm], @{attributes [transfer_domain_rule]})] end end fun get_pcrel_info ctxt qty_full_name = #pcr_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name)) fun get_Domainp_thm quot_thm = the (get_first (try(curry op RS quot_thm)) [@{thm eq_onp_to_Domainp}, @{thm Quotient_to_Domainp}]) fun notes names thms = let val notes = if names then map (fn (name, thms, attrs) => ((name, []), [(thms, attrs)])) thms else map_filter (fn (_, thms, attrs) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) thms in Local_Theory.notes notes #> snd end fun map_thms map_name map_thm thms = map (fn (name, thms, attr) => (map_name name, map map_thm thms, attr)) thms (* Sets up the Lifting package by a quotient theorem. quot_thm - a quotient theorem (Quotient R Abs Rep T) opt_reflp_thm - a theorem saying that a relation from quot_thm is reflexive (in the form "reflp R") opt_par_thm - a parametricity theorem for R *) fun setup_by_quotient (config: config) quot_thm opt_reflp_thm opt_par_thm lthy0 = let (**) val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy0) quot_thm (**) val (rty, qty) = quot_thm_rty_qty quot_thm - val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty)))) + val induct_attr = Attrib.internal \<^here> (K (Induct.induct_type (fst (dest_Type qty)))) val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("abs_induct", @{thms Quotient_total_abs_induct}, [induct_attr]), ("abs_eq_iff", @{thms Quotient_total_abs_eq_iff}, [])] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => let val thms = [("abs_induct", @{thms Quotient_abs_induct}, [induct_attr])] in map_thms qualify (fn thm => quot_thm RS thm) thms end val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val notes2 = map_thms qualify (fn thm => quot_thm RS thm) [("rel_eq_transfer", @{thms Quotient_rel_eq_transfer}, @{attributes [transfer_rule]}), ("right_unique", @{thms Quotient_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms Quotient_right_total}, @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun generate_parametric_rel_eq ctxt transfer_rule opt_param_thm = (case opt_param_thm of NONE => transfer_rule | SOME param_thm => (Lifting_Def.generate_parametric_transfer_rule ctxt transfer_rule param_thm handle Lifting_Term.MERGE_TRANSFER_REL msg => error ("Generation of a parametric transfer rule for the quotient relation failed:\n" ^ Pretty.string_of msg))) fun setup_transfer_rules_par ctxt notes = let val pcrel_info = the (get_pcrel_info ctxt qty_full_name) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val thms = [("id_abs_transfer", [id_abs_transfer], @{attributes [transfer_rule]}), ("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val rel_eq_transfer = generate_parametric_rel_eq ctxt (Lifting_Term.parametrize_transfer_rule ctxt (quot_thm RS @{thm Quotient_rel_eq_transfer})) opt_par_thm val right_unique = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_unique}) val right_total = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_total}) val notes2 = map_thms qualify I [("rel_eq_transfer", [rel_eq_transfer], @{attributes [transfer_rule]}), ("right_unique", [right_unique], @{attributes [transfer_rule]}), ("right_total", [right_total], @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy0 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end (* Sets up the Lifting package by a typedef theorem. gen_code - flag if an abstract type given by typedef_thm should be registred as an abstract type in the code generator typedef_thm - a typedef theorem (type_definition Rep Abs S) *) fun setup_by_typedef_thm config typedef_thm lthy0 = let val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm val (T_def, lthy1) = define_crel config rep_fun lthy0 (**) val T_def = Morphism.thm (Local_Theory.target_morphism lthy1) T_def (**) val quot_thm = case typedef_set of Const (\<^const_name>\top\, _) => [typedef_thm, T_def] MRSL @{thm UNIV_typedef_to_Quotient} | Const (\<^const_name>\Collect\, _) $ Abs (_, _, _) => [typedef_thm, T_def] MRSL @{thm open_typedef_to_Quotient} | _ => [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient} val (rty, qty) = quot_thm_rty_qty quot_thm val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val opt_reflp_thm = case typedef_set of Const (\<^const_name>\top\, _) => SOME ((typedef_thm RS @{thm UNIV_typedef_to_equivp}) RS @{thm equivp_reflp2}) | _ => NONE val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val thms = [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]}), ("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [typedef_thm, T_def] MRSL thm) thms @ notes1 @ notes end fun setup_transfer_rules_par ctxt notes = let val pcrel_info = (the (get_pcrel_info ctxt qty_full_name)) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val thms = [("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]}), ("id_abs_transfer",[id_abs_transfer], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val notes2 = map_thms qualify (fn thm => generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([typedef_thm, T_def] MRSL thm))) [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]})]; val notes3 = map_thms qualify (fn thm => parametrize_class_constraint ctxt pcrel_def ([typedef_thm, T_def] MRSL thm)) [("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique},@{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]})] in notes3 @ notes2 @ notes1 @ notes end val notes1 = [(Binding.prefix_name "Quotient_" qty_name, [quot_thm], [])] fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy1 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy = let val input_thm = singleton (Attrib.eval_thms lthy) xthm val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." fun sanity_check_reflp_thm reflp_thm = let val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." in case reflp_tm of \<^Const_>\reflp_on _ for \<^Const>\top_class.top _\ _\ => () | _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." end fun check_qty qty = if not (is_Type qty) then error "The abstract type must be a type constructor." else () fun setup_quotient () = let val opt_reflp_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_reflp_xthm val _ = if is_some opt_reflp_thm then sanity_check_reflp_thm (the opt_reflp_thm) else () val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm val _ = check_qty (snd (quot_thm_rty_qty input_thm)) in setup_by_quotient default_config input_thm opt_reflp_thm opt_par_thm lthy |> snd end fun setup_typedef () = let val qty = (range_type o fastype_of o hd o get_args 2) input_term val _ = check_qty qty in case opt_reflp_xthm of SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." | NONE => ( case opt_par_xthm of SOME _ => error "The parametricity theorem cannot be specified if the type_definition theorem is used." | NONE => setup_by_typedef_thm default_config input_thm lthy |> snd ) end in case input_term of (Const (\<^const_name>\Quotient\, _) $ _ $ _ $ _ $ _) => setup_quotient () | (Const (\<^const_name>\type_definition\, _) $ _ $ _ $ _) => setup_typedef () | _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." end val _ = Outer_Syntax.local_theory \<^command_keyword>\setup_lifting\ "setup lifting infrastructure" (Parse.thm -- Scan.option Parse.thm -- Scan.option (\<^keyword>\parametric\ |-- Parse.!!! Parse.thm) >> (fn ((xthm, opt_reflp_xthm), opt_par_xthm) => setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm)) (* restoring lifting infrastructure *) local exception PCR_ERROR of Pretty.T list in fun lifting_restore_sanity_check ctxt (qinfo:Lifting_Info.quotient) = let val quot_thm = (#quot_thm qinfo) val _ = quot_thm_sanity_check ctxt quot_thm val pcr_info_err = (case #pcr_info qinfo of SOME pcr => let val pcrel_def = #pcrel_def pcr val pcr_cr_eq = #pcr_cr_eq pcr val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr definiton theorem is not a plain meta equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcrel_def]] val pcr_const_def = head_of def_lhs val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq)) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr_cr equation theorem is not a plain equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] val (pcr_const_eq, eqs) = strip_comb eq_lhs fun is_eq (Const (\<^const_name>\HOL.eq\, _)) = true | is_eq _ = false fun eq_Const (Const (name1, _)) (Const (name2, _)) = (name1 = name2) | eq_Const _ _ = false val all_eqs = if not (forall is_eq eqs) then [Pretty.block [Pretty.str "Arguments of the lhs of the pcr_cr equation theorem are not only equalities:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] else [] val pcr_consts_not_equal = if not (eq_Const pcr_const_def pcr_const_eq) then [Pretty.block [Pretty.str "Parametrized correspondence relation constants in pcr_def and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_def, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_eq]] else [] val crel = quot_thm_crel quot_thm val cr_consts_not_equal = if not (eq_Const crel eq_rhs) then [Pretty.block [Pretty.str "Correspondence relation constants in the Quotient theorem and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt crel, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt eq_rhs]] else [] in all_eqs @ pcr_consts_not_equal @ cr_consts_not_equal end | NONE => []) val errs = pcr_info_err in if null errs then () else raise PCR_ERROR errs end handle PCR_ERROR errs => error (cat_lines (["Sanity check failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end (* Registers the data in qinfo in the Lifting infrastructure. *) fun lifting_restore qinfo ctxt = let val _ = lifting_restore_sanity_check (Context.proof_of ctxt) qinfo val (_, qty) = quot_thm_rty_qty (#quot_thm qinfo) val qty_full_name = (fst o dest_Type) qty val stored_qinfo = Lifting_Info.lookup_quotients (Context.proof_of ctxt) qty_full_name in if is_some (stored_qinfo) andalso not (Lifting_Info.quotient_eq (qinfo, (the stored_qinfo))) then error (Pretty.string_of (Pretty.block [Pretty.str "Lifting is already setup for the type", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Context.proof_of ctxt) qty)])) else Lifting_Info.update_quotients qty_full_name qinfo ctxt end val parse_opt_pcr = Scan.optional (Attrib.thm -- Attrib.thm >> (fn (pcrel_def, pcr_cr_eq) => SOME {pcrel_def = pcrel_def, pcr_cr_eq = pcr_cr_eq})) NONE val lifting_restore_attribute_setup = Attrib.setup \<^binding>\lifting_restore\ ((Attrib.thm -- parse_opt_pcr) >> (fn (quot_thm, opt_pcr) => let val qinfo = { quot_thm = quot_thm, pcr_info = opt_pcr} in Thm.declaration_attribute (K (lifting_restore qinfo)) end)) "restoring lifting infrastructure" val _ = Theory.setup lifting_restore_attribute_setup fun lifting_restore_internal bundle_name ctxt = let val restore_info = Lifting_Info.lookup_restore_data (Context.proof_of ctxt) bundle_name in case restore_info of SOME restore_info => ctxt |> lifting_restore (#quotient restore_info) |> fold_rev Transfer.transfer_raw_add (Item_Net.content (#transfer_rules restore_info)) | NONE => ctxt end val lifting_restore_internal_attribute_setup = Attrib.setup \<^binding>\lifting_restore_internal\ (Scan.lift Parse.string >> (fn name => Thm.declaration_attribute (K (lifting_restore_internal name)))) "restoring lifting infrastructure; internal attribute; not meant to be used directly by regular users" val _ = Theory.setup lifting_restore_internal_attribute_setup (* lifting_forget *) val monotonicity_names = [\<^const_name>\right_unique\, \<^const_name>\left_unique\, \<^const_name>\right_total\, \<^const_name>\left_total\, \<^const_name>\bi_unique\, \<^const_name>\bi_total\] fun fold_transfer_rel f (Const (\<^const_name>\Transfer.Rel\, _) $ rel $ _ $ _) = f rel | fold_transfer_rel f (Const (\<^const_name>\HOL.eq\, _) $ (Const (\<^const_name>\Domainp\, _) $ rel) $ _) = f rel | fold_transfer_rel f (Const (name, _) $ rel) = if member op= monotonicity_names name then f rel else f \<^term>\undefined\ | fold_transfer_rel f _ = f \<^term>\undefined\ fun filter_transfer_rules_by_rel transfer_rel transfer_rules = let val transfer_rel_name = transfer_rel |> dest_Const |> fst; fun has_transfer_rel thm = let val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop in member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name end handle TERM _ => false in filter has_transfer_rel transfer_rules end type restore_data = {quotient : Lifting_Info.quotient, transfer_rules: thm Item_Net.T} fun get_transfer_rel (qinfo : Lifting_Info.quotient) = let fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of in if is_some (#pcr_info qinfo) then get_pcrel (#pcrel_def (the (#pcr_info qinfo))) else quot_thm_crel (#quot_thm qinfo) end fun pointer_of_bundle_name bundle_name ctxt = let val bundle = Bundle.read ctxt bundle_name fun err () = error "The provided bundle is not a lifting bundle" in (case bundle of [(_, [arg_src])] => let val (name, _) = Token.syntax (Scan.lift Parse.string) arg_src ctxt handle ERROR _ => err () in name end | _ => err ()) end fun pointer_of_bundle_binding ctxt binding = Name_Space.full_name (Name_Space.naming_of (Context.Theory (Proof_Context.theory_of ctxt))) binding fun lifting_forget pointer lthy = let fun get_transfer_rules_to_delete qinfo ctxt = let val transfer_rel = get_transfer_rel qinfo in filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw ctxt) end in case Lifting_Info.lookup_restore_data lthy pointer of SOME restore_info => let val qinfo = #quotient restore_info val quot_thm = #quot_thm qinfo val transfer_rules = get_transfer_rules_to_delete qinfo lthy in - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm)) lthy end | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_forget_cmd bundle_name lthy = lifting_forget (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_forget\ "unsetup Lifting and Transfer for the given lifting bundle" (Parse.name_position >> lifting_forget_cmd) (* lifting_update *) fun update_transfer_rules pointer lthy = let fun new_transfer_rules ({ quotient = qinfo, ... }:Lifting_Info.restore_data) lthy = let val transfer_rel = get_transfer_rel qinfo val transfer_rules = filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw lthy) in fn phi => fold_rev (Item_Net.update o Morphism.thm phi) transfer_rules Thm.item_net end in case Lifting_Info.lookup_restore_data lthy pointer of SOME refresh_data => - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer (new_transfer_rules refresh_data lthy phi)) lthy | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_update_cmd bundle_name lthy = update_transfer_rules (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_update\ "add newly introduced transfer rules to a bundle storing the state of Lifting and Transfer" (Parse.name_position >> lifting_update_cmd) end diff --git a/src/HOL/Tools/Quotient/quotient_def.ML b/src/HOL/Tools/Quotient/quotient_def.ML --- a/src/HOL/Tools/Quotient/quotient_def.ML +++ b/src/HOL/Tools/Quotient/quotient_def.ML @@ -1,210 +1,210 @@ (* Title: HOL/Tools/Quotient/quotient_def.ML Author: Cezary Kaliszyk and Christian Urban Definitions for constants on quotient types. *) signature QUOTIENT_DEF = sig val add_quotient_def: ((binding * mixfix) * Attrib.binding) * (term * term) -> thm -> local_theory -> Quotient_Info.quotconsts * local_theory val quotient_def: (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) -> local_theory -> Proof.state val quotient_def_cmd: (binding * string option * mixfix) option * (Attrib.binding * (string * string)) -> local_theory -> Proof.state end; structure Quotient_Def: QUOTIENT_DEF = struct (** Interface and Syntax Setup **) (* Generation of the code certificate from the rsp theorem *) open Lifting_Util infix 0 MRSL (* The ML-interface for a quotient definition takes as argument: - an optional binding and mixfix annotation - attributes - the new constant as term - the rhs of the definition as term - respectfulness theorem for the rhs It stores the qconst_info in the quotconsts data slot. Restriction: At the moment the left- and right-hand side of the definition must be a constant. *) fun error_msg bind str = let val name = Binding.name_of bind val pos = Position.here (Binding.pos_of bind) in error ("Head of quotient_definition " ^ quote str ^ " differs from declaration " ^ name ^ pos) end fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy = let val rty = fastype_of rhs val qty = fastype_of lhs val absrep_trm = Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) val (_, prop') = Local_Defs.cert_def lthy (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val ((qconst, (_ , def)), lthy') = Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy fun qconst_data phi = Quotient_Info.transform_quotconsts phi {qconst = qconst, rconst = rhs, def = def} fun qualify defname suffix = Binding.name suffix |> Binding.qualify true defname val lhs_name = Binding.name_of (#1 var) val rsp_thm_name = qualify lhs_name "rsp" val lthy'' = lthy' - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => (case qconst_data phi of qcinfo as {qconst = Const (c, _), ...} => Quotient_Info.update_quotconsts (c, qcinfo) | _ => I)) |> (snd oo Local_Theory.note) ((rsp_thm_name, @{attributes [quot_respect]}), [rsp_thm]) in (qconst_data Morphism.identity, lthy'') end fun mk_readable_rsp_thm_eq tm ctxt = let val ctm = Thm.cterm_of ctxt tm fun abs_conv2 cv = Conv.abs_conv (Conv.abs_conv (cv o #2) o #2) ctxt fun erase_quants ctxt' ctm' = case (Thm.term_of ctm') of Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Conv.all_conv ctm' | _ => (Conv.binder_conv (erase_quants o #2) ctxt' then_conv Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm' val norm_fun_eq = abs_conv2 erase_quants then_conv Thm.eta_conversion fun simp_arrows_conv ctm = let val unfold_conv = Conv.rewrs_conv [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, @{thm rel_fun_def[THEN eq_reflection]}] val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 in case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => (binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm | _ => Conv.all_conv ctm end val unfold_ret_val_invs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv) val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt val beta_conv = Thm.beta_conversion true val eq_thm = (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm in Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) end fun gen_quotient_def prep_var parse_term (raw_var, (attr, (raw_lhs, raw_rhs))) lthy = let val (opt_var, ctxt) = (case raw_var of NONE => (NONE, lthy) | SOME var => prep_var var lthy |>> SOME) val lhs_constraints = (case opt_var of SOME (_, SOME T, _) => [T] | _ => []) fun prep_term Ts = parse_term ctxt #> fold Type.constraint Ts #> Syntax.check_term ctxt; val lhs = prep_term lhs_constraints raw_lhs val rhs = prep_term [] raw_rhs val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined" val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" val _ = if is_Const rhs then () else warning "The definiens is not a constant" val var = (case opt_var of NONE => (Binding.name lhs_str, NoSyn) | SOME (binding, _, mx) => if Variable.check_name binding = lhs_str then (binding, mx) else error_msg binding lhs_str); fun try_to_prove_refl thm = let val lhs_eq = thm |> Thm.prop_of |> Logic.dest_implies |> fst |> strip_all_body |> try HOLogic.dest_Trueprop in case lhs_eq of SOME (Const (\<^const_name>\HOL.eq\, _) $ _ $ _) => SOME (@{thm refl} RS thm) | SOME _ => (case body_type (fastype_of lhs) of Type (typ_name, _) => \<^try>\ #equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) RS @{thm Equiv_Relations.equivp_reflp} RS thm\ | _ => NONE ) | _ => NONE end val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty) val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs)) val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) fun after_qed thm_list lthy = let val internal_rsp_thm = case thm_list of [] => the maybe_proven_rsp_thm | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm (fn _ => resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN Proof_Context.fact_tac ctxt [thm] 1) in snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy) end in case maybe_proven_rsp_thm of SOME _ => Proof.theorem NONE after_qed [] lthy | NONE => Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy end val quotient_def = gen_quotient_def Proof_Context.cert_var (K I) val quotient_def_cmd = gen_quotient_def Proof_Context.read_var Syntax.parse_term (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\quotient_definition\ "definition for constants over the quotient type" (Scan.option Parse_Spec.constdecl -- Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term -- (\<^keyword>\is\ |-- Parse.term))) >> quotient_def_cmd); end; diff --git a/src/HOL/Tools/Quotient/quotient_type.ML b/src/HOL/Tools/Quotient/quotient_type.ML --- a/src/HOL/Tools/Quotient/quotient_type.ML +++ b/src/HOL/Tools/Quotient/quotient_type.ML @@ -1,355 +1,355 @@ (* Title: HOL/Tools/Quotient/quotient_type.ML Author: Cezary Kaliszyk and Christian Urban Definition of a quotient type. *) signature QUOTIENT_TYPE = sig val add_quotient_type: {overloaded: bool} -> ((string list * binding * mixfix) * (typ * term * bool) * ((binding * binding) option * thm option)) * thm -> local_theory -> Quotient_Info.quotients * local_theory val quotient_type: {overloaded: bool} -> (string list * binding * mixfix) * (typ * term * bool) * ((binding * binding) option * thm option) -> Proof.context -> Proof.state val quotient_type_cmd: {overloaded: bool} -> (((((string list * binding) * mixfix) * string) * (bool * string)) * (binding * binding) option) * (Facts.ref * Token.src list) option -> Proof.context -> Proof.state end; structure Quotient_Type: QUOTIENT_TYPE = struct (*** definition of quotient types ***) val mem_def1 = @{lemma "y \ Collect S \ S y" by simp} val mem_def2 = @{lemma "S y \ y \ Collect S" by simp} (* constructs the term {c. EX (x::rty). rel x x \ c = Collect (rel x)} *) fun typedef_term rel rty lthy = let val [x, c] = [("x", rty), ("c", HOLogic.mk_setT rty)] |> Variable.variant_frees lthy [rel] |> map Free in HOLogic.Collect_const (HOLogic.mk_setT rty) $ (lambda c (HOLogic.exists_const rty $ lambda x (HOLogic.mk_conj (rel $ x $ x, HOLogic.mk_eq (c, HOLogic.Collect_const rty $ (rel $ x)))))) end (* makes the new type definitions and proves non-emptyness *) fun typedef_make overloaded (vs, qty_name, mx, rel, rty) equiv_thm lthy = let fun typedef_tac ctxt = EVERY1 (map (resolve_tac ctxt o single) [@{thm part_equivp_typedef}, equiv_thm]) in Typedef.add_typedef overloaded (qty_name, map (rpair dummyS) vs, mx) (typedef_term rel rty lthy) NONE typedef_tac lthy end (* tactic to prove the quot_type theorem for the new type *) fun typedef_quot_type_tac ctxt equiv_thm ((_, typedef_info): Typedef.info) = let val rep_thm = #Rep typedef_info RS mem_def1 val rep_inv = #Rep_inverse typedef_info val abs_inv = #Abs_inverse typedef_info val rep_inj = #Rep_inject typedef_info in (resolve_tac ctxt @{thms quot_type.intro} THEN' RANGE [ resolve_tac ctxt [equiv_thm], resolve_tac ctxt [rep_thm], resolve_tac ctxt [rep_inv], resolve_tac ctxt [abs_inv] THEN' resolve_tac ctxt [mem_def2] THEN' assume_tac ctxt, resolve_tac ctxt [rep_inj]]) 1 end (* proves the quot_type theorem for the new type *) fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy = let val quot_type_const = Const (\<^const_name>\quot_type\, fastype_of rel --> fastype_of abs --> fastype_of rep --> \<^typ>\bool\) val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep) in Goal.prove lthy [] [] goal (fn {context = ctxt, ...} => typedef_quot_type_tac ctxt equiv_thm typedef_info) end open Lifting_Util infix 0 MRSL fun define_cr_rel equiv_thm abs_fun lthy = let fun force_type_of_rel rel forced_ty = let val thy = Proof_Context.theory_of lthy val rel_ty = (domain_type o fastype_of) rel val ty_inst = Sign.typ_match thy (rel_ty, forced_ty) Vartab.empty in Envir.subst_term_types ty_inst rel end val (rty, qty) = (dest_funT o fastype_of) abs_fun val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0) val Abs_body = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of Const (\<^const_name>\equivp\, _) $ _ => abs_fun_graph | Const (\<^const_name>\part_equivp\, _) $ rel => HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph) | _ => error "unsupported equivalence theorem" ) val def_term = Abs ("x", rty, Abs ("y", qty, Abs_body)); val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val cr_rel_name = Binding.prefix_name "cr_" qty_name val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy val ((_, (_ , def_thm)), lthy'') = Local_Theory.define ((cr_rel_name, NoSyn), ((Thm.def_binding cr_rel_name, []), fixed_def_term)) lthy' in (def_thm, lthy'') end; fun setup_lifting_package quot3_thm equiv_thm opt_par_thm lthy = let val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot3_thm val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy val (rty, qty) = (dest_funT o fastype_of) abs_fun val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name val (reflp_thm, quot_thm) = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of Const (\<^const_name>\equivp\, _) $ _ => (SOME (equiv_thm RS @{thm equivp_reflp2}), [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp}) | Const (\<^const_name>\part_equivp\, _) $ _ => (NONE, [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient}) | _ => error "unsupported equivalence theorem") val config = { notes = true } in lthy' |> Lifting_Setup.setup_by_quotient config quot_thm reflp_thm opt_par_thm |> snd |> (snd oo Local_Theory.note) ((quotient_thm_name, []), [quot_thm]) end fun init_quotient_infr quot_thm equiv_thm opt_par_thm lthy = let val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_thm val (qtyp, rtyp) = (dest_funT o fastype_of) rep val qty_full_name = (fst o dest_Type) qtyp fun quotients phi = Quotient_Info.transform_quotients phi {qtyp = qtyp, rtyp = rtyp, equiv_rel = rel, equiv_thm = equiv_thm, quot_thm = quot_thm} fun abs_rep phi = Quotient_Info.transform_abs_rep phi {abs = abs, rep = rep} in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Quotient_Info.update_quotients (qty_full_name, quotients phi) #> Quotient_Info.update_abs_rep (qty_full_name, abs_rep phi)) |> setup_lifting_package quot_thm equiv_thm opt_par_thm end (* main function for constructing a quotient type *) fun add_quotient_type overloaded (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), equiv_thm) lthy = let val part_equiv = if partial then equiv_thm else equiv_thm RS @{thm equivp_implies_part_equivp} (* generates the typedef *) val ((_, typedef_info), lthy1) = typedef_make overloaded (vs, qty_name, mx, rel, rty) part_equiv lthy (* abs and rep functions from the typedef *) val Abs_ty = #abs_type (#1 typedef_info) val Rep_ty = #rep_type (#1 typedef_info) val Abs_name = #Abs_name (#1 typedef_info) val Rep_name = #Rep_name (#1 typedef_info) val Abs_const = Const (Abs_name, Rep_ty --> Abs_ty) val Rep_const = Const (Rep_name, Abs_ty --> Rep_ty) (* more useful abs and rep definitions *) val abs_const = Const (\<^const_name>\quot_type.abs\, (rty --> rty --> \<^typ>\bool\) --> (Rep_ty --> Abs_ty) --> rty --> Abs_ty) val rep_const = Const (\<^const_name>\quot_type.rep\, (Abs_ty --> Rep_ty) --> Abs_ty --> rty) val abs_trm = abs_const $ rel $ Abs_const val rep_trm = rep_const $ Rep_const val (rep_name, abs_name) = (case opt_morphs of NONE => (Binding.prefix_name "rep_" qty_name, Binding.prefix_name "abs_" qty_name) | SOME morphs => morphs) val ((_, (_, abs_def)), lthy2) = lthy1 |> Local_Theory.define ((abs_name, NoSyn), ((Thm.def_binding abs_name, []), abs_trm)) val ((_, (_, rep_def)), lthy3) = lthy2 |> Local_Theory.define ((rep_name, NoSyn), ((Thm.def_binding rep_name, []), rep_trm)) (* quot_type theorem *) val quot_thm = typedef_quot_type_thm (rel, Abs_const, Rep_const, part_equiv, typedef_info) lthy3 (* quotient theorem *) val quotient_thm_name = Binding.prefix_name "Quotient3_" qty_name val quotient_thm = (quot_thm RS @{thm quot_type.Quotient}) |> fold_rule lthy3 [abs_def, rep_def] (* name equivalence theorem *) val equiv_thm_name = Binding.suffix_name "_equivp" qty_name (* storing the quotients *) val quotients = {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm, quot_thm = quotient_thm} val lthy4 = lthy3 |> init_quotient_infr quotient_thm equiv_thm opt_par_thm |> (snd oo Local_Theory.note) ((equiv_thm_name, if partial then [] else @{attributes [quot_equiv]}), [equiv_thm]) |> (snd oo Local_Theory.note) ((quotient_thm_name, @{attributes [quot_thm]}), [quotient_thm]) in (quotients, lthy4) end (* sanity checks for the quotient type specifications *) fun sanity_check ((vs, qty_name, _), (rty, rel, _), _) = let val rty_tfreesT = map fst (Term.add_tfreesT rty []) val rel_tfrees = map fst (Term.add_tfrees rel []) val rel_frees = map fst (Term.add_frees rel []) val rel_vars = Term.add_vars rel [] val rel_tvars = Term.add_tvars rel [] val qty_str = Binding.print qty_name ^ ": " val illegal_rel_vars = if null rel_vars andalso null rel_tvars then [] else [qty_str ^ "illegal schematic variable(s) in the relation."] val dup_vs = (case duplicates (op =) vs of [] => [] | dups => [qty_str ^ "duplicate type variable(s) on the lhs: " ^ commas_quote dups]) val extra_rty_tfrees = (case subtract (op =) vs rty_tfreesT of [] => [] | extras => [qty_str ^ "extra type variable(s) on the lhs: " ^ commas_quote extras]) val extra_rel_tfrees = (case subtract (op =) vs rel_tfrees of [] => [] | extras => [qty_str ^ "extra type variable(s) in the relation: " ^ commas_quote extras]) val illegal_rel_frees = (case rel_frees of [] => [] | xs => [qty_str ^ "illegal variable(s) in the relation: " ^ commas_quote xs]) val errs = illegal_rel_vars @ dup_vs @ extra_rty_tfrees @ extra_rel_tfrees @ illegal_rel_frees in if null errs then () else error (cat_lines errs) end (* check for existence of map functions *) fun map_check ctxt (_, (rty, _, _), _) = let fun map_check_aux rty warns = (case rty of Type (_, []) => warns | Type (s, _) => if Symtab.defined (Functor.entries ctxt) s then warns else s :: warns | _ => warns) val warns = map_check_aux rty [] in if null warns then () else warning ("No map function defined for " ^ commas warns ^ ". This will cause problems later on.") end (*** interface and syntax setup ***) (* the ML-interface takes a list of tuples consisting of: - the name of the quotient type - its free type variables (first argument) - its mixfix annotation - the type to be quotient - the partial flag (a boolean) - the relation according to which the type is quotient - optional names of morphisms (rep/abs) - flag if code should be generated by Lifting package it opens a proof-state in which one has to show that the relations are equivalence relations *) fun quotient_type overloaded quot lthy = let (* sanity check *) val _ = sanity_check quot val _ = map_check lthy quot fun mk_goal (rty, rel, partial) = let val equivp_ty = ([rty, rty] ---> \<^typ>\bool\) --> \<^typ>\bool\ val const = if partial then \<^const_name>\part_equivp\ else \<^const_name>\equivp\ in HOLogic.mk_Trueprop (Const (const, equivp_ty) $ rel) end val goal = (mk_goal o #2) quot fun after_qed [[thm]] = (snd oo add_quotient_type overloaded) (quot, thm) in Proof.theorem NONE after_qed [[(goal, [])]] lthy end fun quotient_type_cmd overloaded spec lthy = let fun parse_spec ((((((vs, qty_name), mx), rty_str), (partial, rel_str)), opt_morphs), opt_par_xthm) lthy = let val rty = Syntax.read_typ lthy rty_str val tmp_lthy1 = Variable.declare_typ rty lthy val rel = Syntax.parse_term tmp_lthy1 rel_str |> Type.constraint (rty --> rty --> \<^typ>\bool\) |> Syntax.check_term tmp_lthy1 val tmp_lthy2 = Variable.declare_term rel tmp_lthy1 val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm in (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), tmp_lthy2) end val (spec', _) = parse_spec spec lthy in quotient_type overloaded spec' lthy end (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\quotient_type\ "quotient type definitions (require equivalence proofs)" (* FIXME Parse.type_args_constrained and standard treatment of sort constraints *) (Parse_Spec.overloaded -- (Parse.type_args -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- Parse.typ) -- (\<^keyword>\/\ |-- Scan.optional (Parse.reserved "partial" -- \<^keyword>\:\ >> K true) false -- Parse.term) -- Scan.option (\<^keyword>\morphisms\ |-- Parse.!!! (Parse.binding -- Parse.binding)) -- Scan.option (\<^keyword>\parametric\ |-- Parse.!!! Parse.thm)) >> (fn (overloaded, spec) => quotient_type_cmd {overloaded = overloaded} spec)) end diff --git a/src/HOL/Tools/Transfer/transfer_bnf.ML b/src/HOL/Tools/Transfer/transfer_bnf.ML --- a/src/HOL/Tools/Transfer/transfer_bnf.ML +++ b/src/HOL/Tools/Transfer/transfer_bnf.ML @@ -1,281 +1,281 @@ (* Title: HOL/Tools/Transfer/transfer_bnf.ML Author: Ondrej Kuncar, TU Muenchen Setup for Transfer for types that are BNF. *) signature TRANSFER_BNF = sig exception NO_PRED_DATA of unit val base_name_of_bnf: BNF_Def.bnf -> binding val type_name_of_bnf: BNF_Def.bnf -> string val lookup_defined_pred_data: Proof.context -> string -> Transfer.pred_data val bnf_only_type_ctr: (BNF_Def.bnf -> 'a -> 'a) -> BNF_Def.bnf -> 'a -> 'a end structure Transfer_BNF : TRANSFER_BNF = struct open BNF_Util open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar exception NO_PRED_DATA of unit (* util functions *) fun base_name_of_bnf bnf = Binding.name (Binding.name_of (name_of_bnf bnf)) fun mk_Domainp P = let val PT = fastype_of P val argT = hd (binder_types PT) in Const (\<^const_name>\Domainp\, PT --> argT --> HOLogic.boolT) $ P end fun type_name_of_bnf bnf = T_of_bnf bnf |> dest_Type |> fst fun bnf_only_type_ctr f bnf = if is_Type (T_of_bnf bnf) then f bnf else I fun bnf_of_fp_sugar (fp_sugar:fp_sugar) = nth (#bnfs (#fp_res fp_sugar)) (#fp_res_index fp_sugar) fun fp_sugar_only_type_ctr f fp_sugars = (case filter (is_Type o T_of_bnf o bnf_of_fp_sugar) fp_sugars of [] => I | fp_sugars' => f fp_sugars') (* relation constraints - bi_total & co. *) fun mk_relation_constraint name arg = (Const (name, fastype_of arg --> HOLogic.boolT)) $ arg fun side_constraint_tac bnf constr_defs ctxt = let val thms = constr_defs @ map mk_sym [rel_eq_of_bnf bnf, rel_conversep_of_bnf bnf, rel_OO_of_bnf bnf] in SELECT_GOAL (Local_Defs.unfold0_tac ctxt thms) THEN' rtac ctxt (rel_mono_of_bnf bnf) THEN_ALL_NEW assume_tac ctxt end fun bi_constraint_tac constr_iff sided_constr_intros ctxt = SELECT_GOAL (Local_Defs.unfold0_tac ctxt [constr_iff]) THEN' CONJ_WRAP' (fn thm => rtac ctxt thm THEN_ALL_NEW (REPEAT_DETERM o etac ctxt conjE THEN' assume_tac ctxt)) sided_constr_intros fun generate_relation_constraint_goal ctxt bnf constraint_def = let val constr_name = constraint_def |> Thm.prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq |> head_of |> fst o dest_Const val live = live_of_bnf bnf val (((As, Bs), Ds), ctxt') = ctxt |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees' (map Type.sort_of_atyp (deads_of_bnf bnf)) val relator = mk_rel_of_bnf Ds As Bs bnf val relsT = map2 mk_pred2T As Bs val (args, ctxt'') = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt' val concl = HOLogic.mk_Trueprop (mk_relation_constraint constr_name (list_comb (relator, args))) val assms = map (HOLogic.mk_Trueprop o (mk_relation_constraint constr_name)) args val goal = Logic.list_implies (assms, concl) in (goal, ctxt'') end fun prove_relation_side_constraint ctxt bnf constraint_def = let val (goal, ctxt') = generate_relation_constraint_goal ctxt bnf constraint_def in Goal.prove_sorry ctxt' [] [] goal (fn {context = goal_ctxt, ...} => side_constraint_tac bnf [constraint_def] goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt' ctxt) |> Drule.zero_var_indexes end fun prove_relation_bi_constraint ctxt bnf constraint_def side_constraints = let val (goal, ctxt') = generate_relation_constraint_goal ctxt bnf constraint_def in Goal.prove_sorry ctxt' [] [] goal (fn {context = goal_ctxt, ...} => bi_constraint_tac constraint_def side_constraints goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt' ctxt) |> Drule.zero_var_indexes end val defs = [("left_total_rel", @{thm left_total_alt_def}), ("right_total_rel", @{thm right_total_alt_def}), ("left_unique_rel", @{thm left_unique_alt_def}), ("right_unique_rel", @{thm right_unique_alt_def})] fun prove_relation_constraints bnf ctxt = let val transfer_attr = @{attributes [transfer_rule]} val Tname = base_name_of_bnf bnf val defs = map (apsnd (prove_relation_side_constraint ctxt bnf)) defs val bi_total = prove_relation_bi_constraint ctxt bnf @{thm bi_total_alt_def} [snd (nth defs 0), snd (nth defs 1)] val bi_unique = prove_relation_bi_constraint ctxt bnf @{thm bi_unique_alt_def} [snd (nth defs 2), snd (nth defs 3)] val defs = ("bi_total_rel", bi_total) :: ("bi_unique_rel", bi_unique) :: defs in maps (fn (a, thm) => [((Binding.qualify_name true Tname a, []), [([thm], transfer_attr)])]) defs end (* relator_eq *) fun relator_eq bnf = [(Binding.empty_atts, [([rel_eq_of_bnf bnf], @{attributes [relator_eq]})])] (* transfer rules *) fun bnf_transfer_rules bnf = let val transfer_rules = map_transfer_of_bnf bnf :: pred_transfer_of_bnf bnf :: rel_transfer_of_bnf bnf :: set_transfer_of_bnf bnf val transfer_attr = @{attributes [transfer_rule]} in map (fn thm => (Binding.empty_atts, [([thm],transfer_attr)])) transfer_rules end (* Domainp theorem for predicator *) fun Domainp_tac bnf pred_def ctxt = let val n = live_of_bnf bnf val set_map's = set_map_of_bnf bnf in EVERY' [rtac ctxt ext, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [@{thm Domainp.simps}, in_rel_of_bnf bnf, pred_def]), rtac ctxt iffI, REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], hyp_subst_tac ctxt, CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt ballI, dtac ctxt (set_map RS equalityD1 RS set_mp), etac ctxt imageE, dtac ctxt set_rev_mp, assume_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm case_prodE}], hyp_subst_tac ctxt, rtac ctxt @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]}, etac ctxt @{thm DomainPI}]) set_map's, REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM o resolve_tac ctxt [exI, (refl RS conjI), rotate_prems 1 conjI], rtac ctxt refl, rtac ctxt (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym, map_id_of_bnf bnf]), REPEAT_DETERM_N n o (EVERY' [rtac ctxt @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]}, rtac ctxt @{thm fst_conv}]), rtac ctxt CollectI, CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt (set_map RS @{thm ord_eq_le_trans}), REPEAT_DETERM o resolve_tac ctxt [@{thm image_subsetI}, CollectI, @{thm case_prodI}], dtac ctxt (rotate_prems 1 bspec), assume_tac ctxt, etac ctxt @{thm DomainpE}, etac ctxt @{thm someI}]) set_map's ] end fun prove_Domainp_rel ctxt bnf pred_def = let val live = live_of_bnf bnf val (((As, Bs), Ds), ctxt') = ctxt |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees' (map Type.sort_of_atyp (deads_of_bnf bnf)) val relator = mk_rel_of_bnf Ds As Bs bnf val relsT = map2 mk_pred2T As Bs val (args, ctxt'') = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt' val lhs = mk_Domainp (list_comb (relator, args)) val rhs = Term.list_comb (mk_pred_of_bnf Ds As bnf, map mk_Domainp args) val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop in Goal.prove_sorry ctxt'' [] [] goal (fn {context = goal_ctxt, ...} => Domainp_tac bnf pred_def goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt'' ctxt) |> Drule.zero_var_indexes end fun predicator bnf lthy = let val pred_def = pred_set_of_bnf bnf val Domainp_rel = prove_Domainp_rel lthy bnf pred_def val rel_eq_onp = rel_eq_onp_of_bnf bnf val Domainp_rel_thm_name = Binding.qualify_name true (base_name_of_bnf bnf) "Domainp_rel" val pred_data = Transfer.mk_pred_data pred_def rel_eq_onp [] val type_name = type_name_of_bnf bnf val relator_domain_attr = @{attributes [relator_domain]} val notes = [((Domainp_rel_thm_name, []), [([Domainp_rel], relator_domain_attr)])] in lthy |> Local_Theory.notes notes |> snd - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data)) end (* BNF interpretation *) fun transfer_bnf_interpretation bnf lthy = let val dead = dead_of_bnf bnf val constr_notes = if dead > 0 then [] else prove_relation_constraints bnf lthy val relator_eq_notes = if dead > 0 then [] else relator_eq bnf val transfer_rule_notes = if dead > 0 then [] else bnf_transfer_rules bnf in lthy |> Local_Theory.notes (constr_notes @ relator_eq_notes @ transfer_rule_notes) |> snd |> predicator bnf end val _ = Theory.setup (bnf_interpretation transfer_plugin (bnf_only_type_ctr transfer_bnf_interpretation)) (* simplification rules for the predicator *) fun lookup_defined_pred_data ctxt name = case Transfer.lookup_pred_data ctxt name of SOME data => data | NONE => raise NO_PRED_DATA () (* fp_sugar interpretation *) fun fp_sugar_transfer_rules (fp_sugar:fp_sugar) = let val fp_ctr_sugar = #fp_ctr_sugar fp_sugar val transfer_rules = #ctr_transfers fp_ctr_sugar @ #case_transfers fp_ctr_sugar @ #disc_transfers fp_ctr_sugar @ #sel_transfers fp_ctr_sugar @ these (Option.map #co_rec_transfers (#fp_co_induct_sugar fp_sugar)) val transfer_attr = @{attributes [transfer_rule]} in map (fn thm => (Binding.empty_atts, [([thm], transfer_attr)])) transfer_rules end fun register_pred_injects fp_sugar lthy = let val pred_injects = #pred_injects (#fp_bnf_sugar fp_sugar) val type_name = type_name_of_bnf (#fp_bnf fp_sugar) val pred_data = lookup_defined_pred_data lthy type_name |> Transfer.update_pred_simps pred_injects in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data)) end fun transfer_fp_sugars_interpretation fp_sugar lthy = let val lthy = register_pred_injects fp_sugar lthy val transfer_rules_notes = fp_sugar_transfer_rules fp_sugar in lthy |> Local_Theory.notes transfer_rules_notes |> snd end val _ = Theory.setup (fp_sugars_interpretation transfer_plugin (fp_sugar_only_type_ctr (fold transfer_fp_sugars_interpretation))) end diff --git a/src/HOL/Tools/functor.ML b/src/HOL/Tools/functor.ML --- a/src/HOL/Tools/functor.ML +++ b/src/HOL/Tools/functor.ML @@ -1,281 +1,281 @@ (* Title: HOL/Tools/functor.ML Author: Florian Haftmann, TU Muenchen Functorial structure of types. *) signature FUNCTOR = sig val find_atomic: Proof.context -> typ -> (typ * (bool * bool)) list val construct_mapper: Proof.context -> (string * bool -> term) -> bool -> typ -> typ -> term val functor_: string option -> term -> local_theory -> Proof.state val functor_cmd: string option -> string -> Proof.context -> Proof.state type entry val entries: Proof.context -> entry list Symtab.table end; structure Functor : FUNCTOR = struct (* bookkeeping *) val compN = "comp"; val idN = "id"; val compositionalityN = "compositionality"; val identityN = "identity"; type entry = { mapper: term, variances: (sort * (bool * bool)) list, comp: thm, id: thm }; structure Data = Generic_Data ( type T = entry list Symtab.table val empty = Symtab.empty fun merge data = Symtab.merge (K true) data ); val entries = Data.get o Context.Proof; (* type analysis *) fun term_with_typ ctxt T t = Envir.subst_term_types (Sign.typ_match (Proof_Context.theory_of ctxt) (fastype_of t, T) Vartab.empty) t; fun find_atomic ctxt T = let val variances_of = Option.map #variances o try hd o Symtab.lookup_list (entries ctxt); fun add_variance is_contra T = AList.map_default (op =) (T, (false, false)) ((if is_contra then apsnd else apfst) (K true)); fun analyze' is_contra (_, (co, contra)) T = (if co then analyze is_contra T else I) #> (if contra then analyze (not is_contra) T else I) and analyze is_contra (T as Type (tyco, Ts)) = (case variances_of tyco of NONE => add_variance is_contra T | SOME variances => fold2 (analyze' is_contra) variances Ts) | analyze is_contra T = add_variance is_contra T; in analyze false T [] end; fun construct_mapper ctxt atomic = let val lookup = hd o Symtab.lookup_list (entries ctxt); fun constructs is_contra (_, (co, contra)) T T' = (if co then [construct is_contra T T'] else []) @ (if contra then [construct (not is_contra) T T'] else []) and construct is_contra (T as Type (tyco, Ts)) (T' as Type (_, Ts')) = let val { mapper = raw_mapper, variances, ... } = lookup tyco; val args = maps (fn (arg_pattern, (T, T')) => constructs is_contra arg_pattern T T') (variances ~~ (Ts ~~ Ts')); val (U, U') = if is_contra then (T', T) else (T, T'); val mapper = term_with_typ ctxt (map fastype_of args ---> U --> U') raw_mapper; in list_comb (mapper, args) end | construct is_contra (TFree (v, _)) (TFree _) = atomic (v, is_contra); in construct end; (* mapper properties *) val compositionality_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [Simpdata.mk_eq @{thm comp_def}]); fun make_comp_prop ctxt variances (tyco, mapper) = let val sorts = map fst variances val (((vs3, vs2), vs1), _) = ctxt |> Variable.invent_types sorts ||>> Variable.invent_types sorts ||>> Variable.invent_types sorts val (Ts1, Ts2, Ts3) = (map TFree vs1, map TFree vs2, map TFree vs3); fun mk_argT ((T, T'), (_, (co, contra))) = (if co then [(T --> T')] else []) @ (if contra then [(T' --> T)] else []); val contras = maps (fn (_, (co, contra)) => (if co then [false] else []) @ (if contra then [true] else [])) variances; val Ts21 = maps mk_argT ((Ts2 ~~ Ts1) ~~ variances); val Ts32 = maps mk_argT ((Ts3 ~~ Ts2) ~~ variances); fun invents n k nctxt = let val names = Name.invent nctxt n k; in (names, fold Name.declare names nctxt) end; val ((names21, names32), nctxt) = Variable.names_of ctxt |> invents "f" (length Ts21) ||>> invents "f" (length Ts32); val T1 = Type (tyco, Ts1); val T2 = Type (tyco, Ts2); val T3 = Type (tyco, Ts3); val (args21, args32) = (names21 ~~ Ts21, names32 ~~ Ts32); val args31 = map2 (fn is_contra => fn ((f21, T21), (f32, T32)) => if not is_contra then HOLogic.mk_comp (Free (f21, T21), Free (f32, T32)) else HOLogic.mk_comp (Free (f32, T32), Free (f21, T21)) ) contras (args21 ~~ args32) fun mk_mapper T T' args = list_comb (term_with_typ ctxt (map fastype_of args ---> T --> T') mapper, args); val mapper21 = mk_mapper T2 T1 (map Free args21); val mapper32 = mk_mapper T3 T2 (map Free args32); val mapper31 = mk_mapper T3 T1 args31; val eq1 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (HOLogic.mk_comp (mapper21, mapper32), mapper31); val x = Free (the_single (Name.invent nctxt (Long_Name.base_name tyco) 1), T3) val eq2 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (mapper21 $ (mapper32 $ x), mapper31 $ x); val comp_prop = fold_rev Logic.all (map Free (args21 @ args32)) eq1; val compositionality_prop = fold_rev Logic.all (map Free (args21 @ args32) @ [x]) eq2; fun prove_compositionality ctxt comp_thm = Goal.prove_sorry ctxt [] [] compositionality_prop (K (ALLGOALS (Method.insert_tac ctxt [@{thm fun_cong} OF [comp_thm]] THEN' Simplifier.asm_lr_simp_tac (put_simpset compositionality_ss ctxt) THEN_ALL_NEW (Goal.assume_rule_tac ctxt)))); in (comp_prop, prove_compositionality) end; val identity_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [Simpdata.mk_eq @{thm id_def}]); fun make_id_prop ctxt variances (tyco, mapper) = let val (vs, _) = Variable.invent_types (map fst variances) ctxt; val Ts = map TFree vs; fun bool_num b = if b then 1 else 0; fun mk_argT (T, (_, (co, contra))) = replicate (bool_num co + bool_num contra) T val arg_Ts = maps mk_argT (Ts ~~ variances) val T = Type (tyco, Ts); val head = term_with_typ ctxt (map (fn T => T --> T) arg_Ts ---> T --> T) mapper; val lhs1 = list_comb (head, map (HOLogic.id_const) arg_Ts); val lhs2 = list_comb (head, map (fn arg_T => Abs ("x", arg_T, Bound 0)) arg_Ts); val rhs = HOLogic.id_const T; val (id_prop, identity_prop) = apply2 (HOLogic.mk_Trueprop o HOLogic.mk_eq o rpair rhs) (lhs1, lhs2); fun prove_identity ctxt id_thm = Goal.prove_sorry ctxt [] [] identity_prop (K (ALLGOALS (Method.insert_tac ctxt [id_thm] THEN' Simplifier.asm_lr_simp_tac (put_simpset identity_ss ctxt)))); in (id_prop, prove_identity) end; (* analyzing and registering mappers *) fun consume _ _ [] = (false, []) | consume eq x (ys as z :: zs) = if eq (x, z) then (true, zs) else (false, ys); fun split_mapper_typ "fun" T = let val (Ts', T') = strip_type T; val (Ts'', T'') = split_last Ts'; val (Ts''', T''') = split_last Ts''; in (Ts''', T''', T'' --> T') end | split_mapper_typ _ T = let val (Ts', T') = strip_type T; val (Ts'', T'') = split_last Ts'; in (Ts'', T'', T') end; fun analyze_mapper ctxt input_mapper = let val T = fastype_of input_mapper; val _ = Type.no_tvars T; val _ = if null (subtract (op =) (Term.add_tfreesT T []) (Term.add_tfrees input_mapper [])) then () else error ("Illegal additional type variable(s) in term: " ^ Syntax.string_of_term ctxt input_mapper); val _ = if null (Term.add_vars (singleton (Variable.export_terms (Proof_Context.augment input_mapper ctxt) ctxt) input_mapper) []) then () else error ("Illegal locally free variable(s) in term: " ^ Syntax.string_of_term ctxt input_mapper); val mapper = singleton (Variable.polymorphic ctxt) input_mapper; val _ = if null (Term.add_tfreesT (fastype_of mapper) []) then () else error ("Illegal locally fixed type variable(s) in type: " ^ Syntax.string_of_typ ctxt T); fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts | add_tycos _ = I; val tycos = add_tycos T []; val tyco = if tycos = ["fun"] then "fun" else case remove (op =) "fun" tycos of [tyco] => tyco | _ => error ("Bad number of type constructors: " ^ Syntax.string_of_typ ctxt T); in (mapper, T, tyco) end; fun analyze_variances ctxt tyco T = let fun bad_typ () = error ("Bad mapper type: " ^ Syntax.string_of_typ ctxt T); val (Ts, T1, T2) = split_mapper_typ tyco T handle List.Empty => bad_typ (); val _ = apply2 ((fn tyco' => if tyco' = tyco then () else bad_typ ()) o fst o dest_Type) (T1, T2) handle TYPE _ => bad_typ (); val (vs1, vs2) = apply2 (map dest_TFree o snd o dest_Type) (T1, T2) handle TYPE _ => bad_typ (); val _ = if has_duplicates (eq_fst (op =)) (vs1 @ vs2) then bad_typ () else (); fun check_variance_pair (var1 as (_, sort1), var2 as (_, sort2)) = let val coT = TFree var1 --> TFree var2; val contraT = TFree var2 --> TFree var1; val sort = Sign.inter_sort (Proof_Context.theory_of ctxt) (sort1, sort2); in consume (op =) coT ##>> consume (op =) contraT #>> pair sort end; val (variances, left_variances) = fold_map check_variance_pair (vs1 ~~ vs2) Ts; val _ = if null left_variances then () else bad_typ (); in variances end; fun gen_functor prep_term some_prfx raw_mapper lthy = let val (mapper, T, tyco) = analyze_mapper lthy (prep_term lthy raw_mapper); val prfx = the_default (Long_Name.base_name tyco) some_prfx; val variances = analyze_variances lthy tyco T; val (comp_prop, prove_compositionality) = make_comp_prop lthy variances (tyco, mapper); val (id_prop, prove_identity) = make_id_prop lthy variances (tyco, mapper); val qualify = Binding.qualify true prfx o Binding.name; fun mapper_declaration comp_thm id_thm phi context = let val typ_instance = Sign.typ_instance (Context.theory_of context); val mapper' = Morphism.term phi mapper; val T_T' = apply2 fastype_of (mapper, mapper'); val vars = Term.add_vars mapper' []; in if null vars andalso typ_instance T_T' andalso typ_instance (swap T_T') then (Data.map o Symtab.cons_list) (tyco, { mapper = mapper', variances = variances, comp = Morphism.thm phi comp_thm, id = Morphism.thm phi id_thm }) context else context end; fun after_qed [single_comp_thm, single_id_thm] lthy = lthy |> Local_Theory.note ((qualify compN, []), single_comp_thm) ||>> Local_Theory.note ((qualify idN, []), single_id_thm) |-> (fn ((_, [comp_thm]), (_, [id_thm])) => fn lthy => lthy |> Local_Theory.note ((qualify compositionalityN, []), [prove_compositionality lthy comp_thm]) |> snd |> Local_Theory.note ((qualify identityN, []), [prove_identity lthy id_thm]) |> snd - |> Local_Theory.declaration {syntax = false, pervasive = false} + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (mapper_declaration comp_thm id_thm)) in lthy |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) [comp_prop, id_prop]) end val functor_ = gen_functor Syntax.check_term; val functor_cmd = gen_functor Syntax.read_term; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\functor\ "register operations managing the functorial structure of a type" (Scan.option (Parse.name --| \<^keyword>\:\) -- Parse.term >> uncurry functor_cmd); end; diff --git a/src/HOL/Tools/inductive.ML b/src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML +++ b/src/HOL/Tools/inductive.ML @@ -1,1318 +1,1318 @@ (* Title: HOL/Tools/inductive.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer and Markus Wenzel, TU Muenchen (Co)Inductive Definition module for HOL. Features: * least or greatest fixedpoints * mutually recursive definitions * definitions involving arbitrary monotone operators * automatically proves introduction and elimination rules Introduction rules have the form [| M Pj ti, ..., Q x, ... |] ==> Pk t where M is some monotone operator (usually the identity) Q x is any side condition on the free variables ti, t are any terms Pj, Pk are two of the predicates being defined in mutual recursion *) signature INDUCTIVE = sig type result = {preds: term list, elims: thm list, raw_induct: thm, induct: thm, inducts: thm list, intrs: thm list, eqs: thm list} val transform_result: morphism -> result -> result type info = {names: string list, coind: bool} * result val the_inductive: Proof.context -> term -> info val the_inductive_global: Proof.context -> string -> info val print_inductives: bool -> Proof.context -> unit val get_monos: Proof.context -> thm list val mono_add: attribute val mono_del: attribute val mk_cases_tac: Proof.context -> tactic val mk_cases: Proof.context -> term -> thm val inductive_forall_def: thm val rulify: Proof.context -> thm -> thm val inductive_cases: (Attrib.binding * term list) list -> local_theory -> (string * thm list) list * local_theory val inductive_cases_cmd: (Attrib.binding * string list) list -> local_theory -> (string * thm list) list * local_theory val ind_cases_rules: Proof.context -> string list -> (binding * string option * mixfix) list -> thm list val inductive_simps: (Attrib.binding * term list) list -> local_theory -> (string * thm list) list * local_theory val inductive_simps_cmd: (Attrib.binding * string list) list -> local_theory -> (string * thm list) list * local_theory type flags = {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool} val add_inductive: flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> result * local_theory val add_inductive_cmd: bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> result * local_theory val arities_of: thm -> (string * int) list val params_of: thm -> term list val partition_rules: thm -> thm list -> (string * thm list) list val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list val unpartition_rules: thm list -> (string * 'a list) list -> 'a list val infer_intro_vars: theory -> thm -> int -> thm list -> term list list val inductive_internals: bool Config.T val select_disj_tac: Proof.context -> int -> int -> int -> tactic type add_ind_def = flags -> term list -> (Attrib.binding * term) list -> thm list -> term list -> (binding * mixfix) list -> local_theory -> result * local_theory val declare_rules: binding -> bool -> bool -> binding -> string list -> term list -> thm list -> binding list -> Token.src list list -> (thm * string list * int) list -> thm list -> thm -> local_theory -> thm list * thm list * thm list * thm * thm list * local_theory val add_ind_def: add_ind_def val gen_add_inductive: add_ind_def -> flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> result * local_theory val gen_add_inductive_cmd: add_ind_def -> bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> result * local_theory val gen_ind_decl: add_ind_def -> bool -> (local_theory -> local_theory) parser end; structure Inductive: INDUCTIVE = struct (** theory context references **) val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_conj_def = @{thm HOL.induct_conj_def}; val inductive_conj = @{thms induct_conj}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; val inductive_rulify_fallback = @{thms induct_rulify_fallback}; val simp_thms1 = map mk_meta_eq @{lemma "(\ True) = False" "(\ False) = True" "(True \ P) = P" "(False \ P) = True" "(P \ True) = P" "(True \ P) = P" by (fact simp_thms)+}; val simp_thms2 = map mk_meta_eq [@{thm inf_fun_def}, @{thm inf_bool_def}] @ simp_thms1; val simp_thms3 = @{thms le_rel_bool_arg_iff if_False if_True conj_ac le_fun_def le_bool_def sup_fun_def sup_bool_def simp_thms if_bool_eq_disj all_simps ex_simps imp_conjL}; (** misc utilities **) val inductive_internals = Attrib.setup_config_bool \<^binding>\inductive_internals\ (K false); fun message quiet_mode s = if quiet_mode then () else writeln s; fun clean_message ctxt quiet_mode s = if Config.get ctxt quick_and_dirty then () else message quiet_mode s; fun coind_prefix true = "co" | coind_prefix false = ""; fun log (b: int) m n = if m >= n then 0 else 1 + log b (b * m) n; fun make_bool_args f g [] i = [] | make_bool_args f g (x :: xs) i = (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2); fun make_bool_args' xs = make_bool_args (K \<^term>\False\) (K \<^term>\True\) xs; fun arg_types_of k c = drop k (binder_types (fastype_of c)); fun find_arg T x [] = raise Fail "find_arg" | find_arg T x ((p as (_, (SOME _, _))) :: ps) = apsnd (cons p) (find_arg T x ps) | find_arg T x ((p as (U, (NONE, y))) :: ps) = if (T: typ) = U then (y, (U, (SOME x, y)) :: ps) else apsnd (cons p) (find_arg T x ps); fun make_args Ts xs = map (fn (T, (NONE, ())) => Const (\<^const_name>\undefined\, T) | (_, (SOME t, ())) => t) (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts)); fun make_args' Ts xs Us = fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs)); fun dest_predicate cs params t = let val k = length params; val (c, ts) = strip_comb t; val (xs, ys) = chop k ts; val i = find_index (fn c' => c' = c) cs; in if xs = params andalso i >= 0 then SOME (c, i, ys, chop (length ys) (arg_types_of k c)) else NONE end; fun mk_names a 0 = [] | mk_names a 1 = [a] | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n); fun select_disj_tac ctxt = let fun tacs 1 1 = [] | tacs _ 1 = [resolve_tac ctxt @{thms disjI1}] | tacs n i = resolve_tac ctxt @{thms disjI2} :: tacs (n - 1) (i - 1); in fn n => fn i => EVERY' (tacs n i) end; (** context data **) type result = {preds: term list, elims: thm list, raw_induct: thm, induct: thm, inducts: thm list, intrs: thm list, eqs: thm list}; fun transform_result phi {preds, elims, raw_induct: thm, induct, inducts, intrs, eqs} = let val term = Morphism.term phi; val thm = Morphism.thm phi; val fact = Morphism.fact phi; in {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct, induct = thm induct, inducts = fact inducts, intrs = fact intrs, eqs = fact eqs} end; type info = {names: string list, coind: bool} * result; val empty_infos = Item_Net.init (op = o apply2 (#names o fst)) (#preds o snd) val empty_equations = Item_Net.init Thm.eq_thm_prop (single o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of); datatype data = Data of {infos: info Item_Net.T, monos: thm list, equations: thm Item_Net.T}; fun make_data (infos, monos, equations) = Data {infos = infos, monos = monos, equations = equations}; structure Data = Generic_Data ( type T = data; val empty = make_data (empty_infos, [], empty_equations); fun merge (Data {infos = infos1, monos = monos1, equations = equations1}, Data {infos = infos2, monos = monos2, equations = equations2}) = make_data (Item_Net.merge (infos1, infos2), Thm.merge_thms (monos1, monos2), Item_Net.merge (equations1, equations2)); ); fun map_data f = Data.map (fn Data {infos, monos, equations} => make_data (f (infos, monos, equations))); fun rep_data ctxt = Data.get (Context.Proof ctxt) |> (fn Data rep => rep); fun print_inductives verbose ctxt = let val {infos, monos, ...} = rep_data ctxt; val space = Consts.space_of (Proof_Context.consts_of ctxt); val consts = Item_Net.content infos |> maps (fn ({names, ...}, result) => map (rpair result) names) in [Pretty.block (Pretty.breaks (Pretty.str "(co)inductives:" :: map (Pretty.mark_str o #1) (Name_Space.markup_entries verbose ctxt space consts))), Pretty.big_list "monotonicity rules:" (map (Thm.pretty_thm_item ctxt) monos)] end |> Pretty.writeln_chunks; (* inductive info *) fun the_inductive ctxt term = Item_Net.retrieve (#infos (rep_data ctxt)) term |> the_single |> apsnd (transform_result (Morphism.transfer_morphism' ctxt)) fun the_inductive_global ctxt name = #infos (rep_data ctxt) |> Item_Net.content |> filter (fn ({names, ...}, _) => member op = names name) |> the_single |> apsnd (transform_result (Morphism.transfer_morphism' ctxt)) fun put_inductives info = map_data (fn (infos, monos, equations) => (Item_Net.update (apsnd (transform_result Morphism.trim_context_morphism) info) infos, monos, equations)); (* monotonicity rules *) fun get_monos ctxt = #monos (rep_data ctxt) |> map (Thm.transfer' ctxt); fun mk_mono ctxt thm = let fun eq_to_mono thm' = thm' RS (thm' RS @{thm eq_to_mono}); fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD}) handle THM _ => thm RS @{thm le_boolD} in (case Thm.concl_of thm of Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => eq_to_mono (HOLogic.mk_obj_eq thm) | _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _) => eq_to_mono thm | _ $ (Const (\<^const_name>\Orderings.less_eq\, _) $ _ $ _) => dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}])) thm)) | _ => thm) end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Thm.string_of_thm ctxt thm); val mono_add = Thm.declaration_attribute (fn thm => fn context => map_data (fn (infos, monos, equations) => (infos, Thm.add_thm (Thm.trim_context (mk_mono (Context.proof_of context) thm)) monos, equations)) context); val mono_del = Thm.declaration_attribute (fn thm => fn context => map_data (fn (infos, monos, equations) => (infos, Thm.del_thm (mk_mono (Context.proof_of context) thm) monos, equations)) context); val _ = Theory.setup (Attrib.setup \<^binding>\mono\ (Attrib.add_del mono_add mono_del) "declaration of monotonicity rule"); (* equations *) fun retrieve_equations ctxt = Item_Net.retrieve (#equations (rep_data ctxt)) #> map (Thm.transfer' ctxt); val equation_add_permissive = Thm.declaration_attribute (fn thm => map_data (fn (infos, monos, equations) => (infos, monos, perhaps (try (Item_Net.update (Thm.trim_context thm))) equations))); (** process rules **) local fun err_in_rule ctxt name t msg = error (cat_lines ["Ill-formed introduction rule " ^ Binding.print name, Syntax.string_of_term ctxt t, msg]); fun err_in_prem ctxt name t p msg = error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p, "in introduction rule " ^ Binding.print name, Syntax.string_of_term ctxt t, msg]); val bad_concl = "Conclusion of introduction rule must be an inductive predicate"; val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate"; val bad_app = "Inductive predicate must be applied to parameter(s) "; fun atomize_term thy = Raw_Simplifier.rewrite_term thy inductive_atomize []; in fun check_rule ctxt cs params ((binding, att), rule) = let val params' = Term.variant_frees rule (Logic.strip_params rule); val frees = rev (map Free params'); val concl = subst_bounds (frees, Logic.strip_assums_concl rule); val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule); val rule' = Logic.list_implies (prems, concl); val aprems = map (atomize_term (Proof_Context.theory_of ctxt)) prems; val arule = fold_rev (Logic.all o Free) params' (Logic.list_implies (aprems, concl)); fun check_ind err t = (case dest_predicate cs params t of NONE => err (bad_app ^ commas (map (Syntax.string_of_term ctxt) params)) | SOME (_, _, ys, _) => if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs then err bad_ind_occ else ()); fun check_prem' prem t = if member (op =) cs (head_of t) then check_ind (err_in_prem ctxt binding rule prem) t else (case t of Abs (_, _, t) => check_prem' prem t | t $ u => (check_prem' prem t; check_prem' prem u) | _ => ()); fun check_prem (prem, aprem) = if can HOLogic.dest_Trueprop aprem then check_prem' prem prem else err_in_prem ctxt binding rule prem "Non-atomic premise"; val _ = (case concl of Const (\<^const_name>\Trueprop\, _) $ t => if member (op =) cs (head_of t) then (check_ind (err_in_rule ctxt binding rule') t; List.app check_prem (prems ~~ aprems)) else err_in_rule ctxt binding rule' bad_concl | _ => err_in_rule ctxt binding rule' bad_concl); in ((binding, att), arule) end; fun rulify ctxt = hol_simplify ctxt inductive_conj #> hol_simplify ctxt inductive_rulify #> hol_simplify ctxt inductive_rulify_fallback #> Simplifier.norm_hhf ctxt; end; (** proofs for (co)inductive predicates **) (* prove monotonicity *) fun prove_mono quiet_mode skip_mono predT fp_fun monos ctxt = (message (quiet_mode orelse skip_mono andalso Config.get ctxt quick_and_dirty) " Proving monotonicity ..."; (if skip_mono then Goal.prove_sorry else Goal.prove_future) ctxt [] [] (HOLogic.mk_Trueprop (\<^Const>\monotone_on predT predT for \<^Const>\top \<^Type>\set predT\\ \<^Const>\less_eq predT\ \<^Const>\less_eq predT\ fp_fun\)) (fn _ => EVERY [resolve_tac ctxt @{thms monoI} 1, REPEAT (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}] 1), REPEAT (FIRST [assume_tac ctxt 1, resolve_tac ctxt (map (mk_mono ctxt) monos @ get_monos ctxt) 1, eresolve_tac ctxt @{thms le_funE} 1, dresolve_tac ctxt @{thms le_boolD} 1])])); (* prove introduction rules *) fun prove_intrs quiet_mode coind mono fp_def k intr_ts rec_preds_defs ctxt ctxt' = let val _ = clean_message ctxt quiet_mode " Proving the introduction rules ..."; val unfold = funpow k (fn th => th RS fun_cong) (mono RS (fp_def RS (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold}))); val rules = [refl, TrueI, @{lemma "\ False" by (rule notI)}, exI, conjI]; val intrs = map_index (fn (i, intr) => Goal.prove_sorry ctxt [] [] intr (fn _ => EVERY [rewrite_goals_tac ctxt rec_preds_defs, resolve_tac ctxt [unfold RS iffD2] 1, select_disj_tac ctxt (length intr_ts) (i + 1) 1, (*Not ares_tac, since refl must be tried before any equality assumptions; backtracking may occur if the premises have extra variables!*) DEPTH_SOLVE_1 (resolve_tac ctxt rules 1 APPEND assume_tac ctxt 1)]) |> singleton (Proof_Context.export ctxt ctxt')) intr_ts in (intrs, unfold) end; (* prove elimination rules *) fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt ctxt''' = let val _ = clean_message ctxt quiet_mode " Proving the elimination rules ..."; val ([pname], ctxt') = Variable.variant_fixes ["P"] ctxt; val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT)); fun dest_intr r = (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))), Logic.strip_assums_hyp r, Logic.strip_params r); val intrs = map dest_intr intr_ts ~~ intr_names; val rules1 = [disjE, exE, FalseE]; val rules2 = [conjE, FalseE, @{lemma "\ True \ R" by (rule notE [OF _ TrueI])}]; fun prove_elim c = let val Ts = arg_types_of (length params) c; val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt'; val frees = map Free (anames ~~ Ts); fun mk_elim_prem ((_, _, us, _), ts, params') = Logic.list_all (params', Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (frees ~~ us) @ ts, P)); val c_intrs = filter (equal c o #1 o #1 o #1) intrs; val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) :: map mk_elim_prem (map #1 c_intrs) in (Goal.prove_sorry ctxt'' [] prems P (fn {context = ctxt4, prems} => EVERY [cut_tac (hd prems) 1, rewrite_goals_tac ctxt4 rec_preds_defs, dresolve_tac ctxt4 [unfold RS iffD1] 1, REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules1)), REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules2)), EVERY (map (fn prem => DEPTH_SOLVE_1 (assume_tac ctxt4 1 ORELSE resolve_tac ctxt [rewrite_rule ctxt4 rec_preds_defs prem, conjI] 1)) (tl prems))]) |> singleton (Proof_Context.export ctxt'' ctxt'''), map #2 c_intrs, length Ts) end in map prove_elim cs end; (* prove simplification equations *) fun prove_eqs quiet_mode cs params intr_ts intrs (elims: (thm * bstring list * int) list) ctxt ctxt'' = (* FIXME ctxt'' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the simplification rules ..."; fun dest_intr r = (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))), Logic.strip_assums_hyp r, Logic.strip_params r); val intr_ts' = map dest_intr intr_ts; fun prove_eq c (elim: thm * 'a * 'b) = let val Ts = arg_types_of (length params) c; val (anames, ctxt') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt; val frees = map Free (anames ~~ Ts); val c_intrs = filter (equal c o #1 o #1 o #1) (intr_ts' ~~ intrs); fun mk_intr_conj (((_, _, us, _), ts, params'), _) = let fun list_ex ([], t) = t | list_ex ((a, T) :: vars, t) = HOLogic.exists_const T $ Abs (a, T, list_ex (vars, t)); val conjs = map2 (curry HOLogic.mk_eq) frees us @ map HOLogic.dest_Trueprop ts; in list_ex (params', if null conjs then \<^term>\True\ else foldr1 HOLogic.mk_conj conjs) end; val lhs = list_comb (c, params @ frees); val rhs = if null c_intrs then \<^term>\False\ else foldr1 HOLogic.mk_disj (map mk_intr_conj c_intrs); val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {context = ctxt'', params, prems, ...} => select_disj_tac ctxt'' (length c_intrs) (i + 1) 1 THEN EVERY (replicate (length params) (resolve_tac ctxt'' @{thms exI} 1)) THEN (if null prems then resolve_tac ctxt'' @{thms TrueI} 1 else let val (prems', last_prem) = split_last prems; in EVERY (map (fn prem => (resolve_tac ctxt'' @{thms conjI} 1 THEN resolve_tac ctxt'' [prem] 1)) prems') THEN resolve_tac ctxt'' [last_prem] 1 end)) ctxt' 1; fun prove_intr2 (((_, _, us, _), ts, params'), intr) = EVERY (replicate (length params') (eresolve_tac ctxt' @{thms exE} 1)) THEN (if null ts andalso null us then resolve_tac ctxt' [intr] 1 else EVERY (replicate (length ts + length us - 1) (eresolve_tac ctxt' @{thms conjE} 1)) THEN Subgoal.FOCUS_PREMS (fn {context = ctxt'', prems, ...} => let val (eqs, prems') = chop (length us) prems; val rew_thms = map (fn th => th RS @{thm eq_reflection}) eqs; in rewrite_goal_tac ctxt'' rew_thms 1 THEN resolve_tac ctxt'' [intr] 1 THEN EVERY (map (fn p => resolve_tac ctxt'' [p] 1) prems') end) ctxt' 1); in Goal.prove_sorry ctxt' [] [] eq (fn _ => resolve_tac ctxt' @{thms iffI} 1 THEN eresolve_tac ctxt' [#1 elim] 1 THEN EVERY (map_index prove_intr1 c_intrs) THEN (if null c_intrs then eresolve_tac ctxt' @{thms FalseE} 1 else let val (c_intrs', last_c_intr) = split_last c_intrs in EVERY (map (fn ci => eresolve_tac ctxt' @{thms disjE} 1 THEN prove_intr2 ci) c_intrs') THEN prove_intr2 last_c_intr end)) |> rulify ctxt' |> singleton (Proof_Context.export ctxt' ctxt'') end; in map2 prove_eq cs elims end; (* derivation of simplified elimination rules *) local (*delete needless equality assumptions*) val refl_thin = Goal.prove_global \<^theory>\HOL\ [] [] \<^prop>\\P. a = a \ P \ P\ (fn {context = ctxt, ...} => assume_tac ctxt 1); val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE]; fun elim_tac ctxt = REPEAT o eresolve_tac ctxt elim_rls; fun simp_case_tac ctxt i = EVERY' [elim_tac ctxt, asm_full_simp_tac ctxt, elim_tac ctxt, REPEAT o bound_hyp_subst_tac ctxt] i; in fun mk_cases_tac ctxt = ALLGOALS (simp_case_tac ctxt) THEN prune_params_tac ctxt; fun mk_cases ctxt prop = let fun err msg = error (Pretty.string_of (Pretty.block [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop])); val elims = Induct.find_casesP ctxt prop; val cprop = Thm.cterm_of ctxt prop; fun mk_elim rl = Thm.implies_intr cprop (Tactic.rule_by_tactic ctxt (mk_cases_tac ctxt) (Thm.assume cprop RS rl)) |> singleton (Proof_Context.export (Proof_Context.augment prop ctxt) ctxt); in (case get_first (try mk_elim) elims of SOME r => r | NONE => err "Proposition not an inductive predicate:") end; end; (* inductive_cases *) fun gen_inductive_cases prep_att prep_prop args lthy = let val thmss = map snd args |> burrow (grouped 10 Par_List.map_independent (mk_cases lthy o prep_prop lthy)); val facts = map2 (fn ((a, atts), _) => fn thms => ((a, map (prep_att lthy) atts), [(thms, [])])) args thmss; in lthy |> Local_Theory.notes facts end; val inductive_cases = gen_inductive_cases (K I) Syntax.check_prop; val inductive_cases_cmd = gen_inductive_cases Attrib.check_src Syntax.read_prop; (* ind_cases *) fun ind_cases_rules ctxt raw_props raw_fixes = let val (props, ctxt') = Specification.read_props raw_props raw_fixes ctxt; val rules = Proof_Context.export ctxt' ctxt (map (mk_cases ctxt') props); in rules end; val _ = Theory.setup (Method.setup \<^binding>\ind_cases\ (Scan.lift (Scan.repeat1 Parse.prop -- Parse.for_fixes) >> (fn (props, fixes) => fn ctxt => Method.erule ctxt 0 (ind_cases_rules ctxt props fixes))) "case analysis for inductive definitions, based on simplified elimination rule"); (* derivation of simplified equation *) fun mk_simp_eq ctxt prop = let val thy = Proof_Context.theory_of ctxt; val ctxt' = Proof_Context.augment prop ctxt; val lhs_of = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of; val substs = retrieve_equations ctxt (HOLogic.dest_Trueprop prop) |> map_filter (fn eq => SOME (Pattern.match thy (lhs_of eq, HOLogic.dest_Trueprop prop) (Vartab.empty, Vartab.empty), eq) handle Pattern.MATCH => NONE); val (subst, eq) = (case substs of [s] => s | _ => error ("equations matching pattern " ^ Syntax.string_of_term ctxt prop ^ " is not unique")); val inst = map (fn v => (fst v, Thm.cterm_of ctxt' (Envir.subst_term subst (Var v)))) (Term.add_vars (lhs_of eq) []); in infer_instantiate ctxt' inst eq |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.full_rewrite ctxt'))) |> singleton (Proof_Context.export ctxt' ctxt) end (* inductive simps *) fun gen_inductive_simps prep_att prep_prop args lthy = let val facts = args |> map (fn ((a, atts), props) => ((a, map (prep_att lthy) atts), map (Thm.no_attributes o single o mk_simp_eq lthy o prep_prop lthy) props)); in lthy |> Local_Theory.notes facts end; val inductive_simps = gen_inductive_simps (K I) Syntax.check_prop; val inductive_simps_cmd = gen_inductive_simps Attrib.check_src Syntax.read_prop; (* prove induction rule *) fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def rec_preds_defs ctxt ctxt''' = (* FIXME ctxt''' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the induction rule ..."; (* predicates for induction rule *) val (pnames, ctxt') = Variable.variant_fixes (mk_names "P" (length cs)) ctxt; val preds = map2 (curry Free) pnames (map (fn c => arg_types_of (length params) c ---> HOLogic.boolT) cs); (* transform an introduction rule into a premise for induction rule *) fun mk_ind_prem r = let fun subst s = (case dest_predicate cs params s of SOME (_, i, ys, (_, Ts)) => let val k = length Ts; val bs = map Bound (k - 1 downto 0); val P = list_comb (nth preds i, map (incr_boundvars k) ys @ bs); val Q = fold_rev Term.abs (mk_names "x" k ~~ Ts) (HOLogic.mk_binop \<^const_name>\HOL.induct_conj\ (list_comb (incr_boundvars k s, bs), P)); in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end | NONE => (case s of t $ u => (fst (subst t) $ fst (subst u), NONE) | Abs (a, T, t) => (Abs (a, T, fst (subst t)), NONE) | _ => (s, NONE))); fun mk_prem s prems = (case subst s of (_, SOME (t, u)) => t :: u :: prems | (t, _) => t :: prems); val SOME (_, i, ys, _) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); in fold_rev (Logic.all o Free) (Logic.strip_params r) (Logic.list_implies (map HOLogic.mk_Trueprop (fold_rev mk_prem (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r)) []), HOLogic.mk_Trueprop (list_comb (nth preds i, ys)))) end; val ind_prems = map mk_ind_prem intr_ts; (* make conclusions for induction rules *) val Tss = map (binder_types o fastype_of) preds; val (xnames, ctxt'') = Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt'; val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (((xnames, Ts), c), P) => let val frees = map Free (xnames ~~ Ts) in HOLogic.mk_imp (list_comb (c, params @ frees), list_comb (P, frees)) end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds))); (* make predicate for instantiation of abstract induction rule *) val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj (map_index (fn (i, P) => fold_rev (curry HOLogic.mk_imp) (make_bool_args HOLogic.mk_not I bs i) (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))) preds)); val ind_concl = HOLogic.mk_Trueprop (HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ (rec_const, ind_pred)); val raw_fp_induct = mono RS (fp_def RS @{thm def_lfp_induct}); val induct = Goal.prove_sorry ctxt'' [] ind_prems ind_concl (fn {context = ctxt3, prems} => EVERY [rewrite_goals_tac ctxt3 [inductive_conj_def], DETERM (resolve_tac ctxt3 [raw_fp_induct] 1), REPEAT (resolve_tac ctxt3 [@{thm le_funI}, @{thm le_boolI}] 1), rewrite_goals_tac ctxt3 simp_thms2, (*This disjE separates out the introduction rules*) REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [disjE, exE, FalseE])), (*Now break down the individual cases. No disjE here in case some premise involves disjunction.*) REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [conjE] ORELSE' bound_hyp_subst_tac ctxt3)), REPEAT (FIRSTGOAL (resolve_tac ctxt3 [conjI, impI] ORELSE' (eresolve_tac ctxt3 [notE] THEN' assume_tac ctxt3))), EVERY (map (fn prem => DEPTH_SOLVE_1 (assume_tac ctxt3 1 ORELSE resolve_tac ctxt3 [rewrite_rule ctxt3 (inductive_conj_def :: rec_preds_defs @ simp_thms2) prem, conjI, refl] 1)) prems)]); val lemma = Goal.prove_sorry ctxt'' [] [] (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn {context = ctxt3, ...} => EVERY [rewrite_goals_tac ctxt3 rec_preds_defs, REPEAT (EVERY [REPEAT (resolve_tac ctxt3 [conjI, impI] 1), REPEAT (eresolve_tac ctxt3 [@{thm le_funE}, @{thm le_boolE}] 1), assume_tac ctxt3 1, rewrite_goals_tac ctxt3 simp_thms1, assume_tac ctxt3 1])]); in singleton (Proof_Context.export ctxt'' ctxt''') (induct RS lemma) end; (* prove coinduction rule *) fun If_const T = Const (\<^const_name>\If\, HOLogic.boolT --> T --> T --> T); fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end; fun prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono fp_def rec_preds_defs ctxt ctxt''' = (* FIXME ctxt''' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the coinduction rule ..."; val n = length cs; val (ns, xss) = map_split (fn pred => make_args' argTs xs (arg_types_of (length params) pred) |> `length) preds; val xTss = map (map fastype_of) xss; val (Rs_names, names_ctxt) = Variable.variant_fixes (mk_names "X" n) ctxt; val Rs = map2 (fn name => fn Ts => Free (name, Ts ---> \<^typ>\bool\)) Rs_names xTss; val Rs_applied = map2 (curry list_comb) Rs xss; val preds_applied = map2 (curry list_comb) (map (fn p => list_comb (p, params)) preds) xss; val abstract_list = fold_rev (absfree o dest_Free); val bss = map (make_bool_args (fn b => HOLogic.mk_eq (b, \<^term>\False\)) (fn b => HOLogic.mk_eq (b, \<^term>\True\)) bs) (0 upto n - 1); val eq_undefinedss = map (fn ys => map (fn x => HOLogic.mk_eq (x, Const (\<^const_name>\undefined\, fastype_of x))) (subtract (op =) ys xs)) xss; val R = @{fold 3} (fn bs => fn eqs => fn R => fn t => if null bs andalso null eqs then R else mk_If (Library.foldr1 HOLogic.mk_conj (bs @ eqs)) R t) bss eq_undefinedss Rs_applied \<^term>\False\ |> abstract_list (bs @ xs); fun subst t = (case dest_predicate cs params t of SOME (_, i, ts, (_, Us)) => let val l = length Us; val bs = map Bound (l - 1 downto 0); val args = map (incr_boundvars l) ts @ bs in HOLogic.mk_disj (list_comb (nth Rs i, args), list_comb (nth preds i, params @ args)) |> fold_rev absdummy Us end | NONE => (case t of t1 $ t2 => subst t1 $ subst t2 | Abs (x, T, u) => Abs (x, T, subst u) | _ => t)); fun mk_coind_prem r = let val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); val ps = map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @ map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r); in (i, fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps)) end; fun mk_prem i Ps = Logic.mk_implies ((nth Rs_applied i, Library.foldr1 HOLogic.mk_disj Ps) |> @{apply 2} HOLogic.mk_Trueprop) |> fold_rev Logic.all (nth xss i); val prems = map mk_coind_prem intr_ts |> AList.group (op =) |> sort (int_ord o apply2 fst) |> map (uncurry mk_prem); val concl = @{map 3} (fn xs => Ctr_Sugar_Util.list_all_free xs oo curry HOLogic.mk_imp) xss Rs_applied preds_applied |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val pred_defs_sym = if null rec_preds_defs then [] else map2 (fn n => fn thm => funpow n (fn thm => thm RS @{thm meta_fun_cong}) thm RS @{thm Pure.symmetric}) ns rec_preds_defs; val simps = simp_thms3 @ pred_defs_sym; val simprocs = [Simplifier.the_simproc ctxt "HOL.defined_All"]; val simplify = asm_full_simplify (Ctr_Sugar_Util.ss_only simps ctxt addsimprocs simprocs); val coind = (mono RS (fp_def RS @{thm def_coinduct})) |> infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt R)] |> simplify; fun idx_of t = find_index (fn R => R = the_single (subtract (op =) (preds @ params) (map Free (Term.add_frees t [])))) Rs; val coind_concls = HOLogic.dest_Trueprop (Thm.concl_of coind) |> HOLogic.dest_conj |> map (fn t => (idx_of t, t)) |> sort (int_ord o @{apply 2} fst) |> map snd; val reorder_bound_goals = map_filter (fn (t, u) => if t aconv u then NONE else SOME (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)))) ((HOLogic.dest_Trueprop concl |> HOLogic.dest_conj) ~~ coind_concls); val reorder_bound_thms = map (fn goal => Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (EVERY' [resolve_tac ctxt [iffI], REPEAT_DETERM o resolve_tac ctxt [allI, impI], REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt, REPEAT_DETERM o resolve_tac ctxt [allI, impI], REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt]))) reorder_bound_goals; val coinduction = Goal.prove_sorry ctxt [] prems concl (fn {context = ctxt, prems = CIH} => HEADGOAL (full_simp_tac (Ctr_Sugar_Util.ss_only (simps @ reorder_bound_thms) ctxt addsimprocs simprocs) THEN' resolve_tac ctxt [coind]) THEN ALLGOALS (REPEAT_ALL_NEW (REPEAT_DETERM o resolve_tac ctxt [allI, impI, conjI] THEN' REPEAT_DETERM o eresolve_tac ctxt [exE, conjE] THEN' dresolve_tac ctxt (map simplify CIH) THEN' REPEAT_DETERM o (assume_tac ctxt ORELSE' eresolve_tac ctxt [conjE] ORELSE' dresolve_tac ctxt [spec, mp])))) in coinduction |> length cs = 1 ? (Object_Logic.rulify ctxt #> rotate_prems ~1) |> singleton (Proof_Context.export names_ctxt ctxt''') end (** specification of (co)inductive predicates **) fun mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy = let val fp_name = if coind then \<^const_name>\Inductive.gfp\ else \<^const_name>\Inductive.lfp\; val argTs = fold (combine (op =) o arg_types_of (length params)) cs []; val k = log 2 1 (length cs); val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT; val p :: xs = map Free (Variable.variant_frees lthy intr_ts (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs))); val bs = map Free (Variable.variant_frees lthy (p :: xs @ intr_ts) (map (rpair HOLogic.boolT) (mk_names "b" k))); fun subst t = (case dest_predicate cs params t of SOME (_, i, ts, (Ts, Us)) => let val l = length Us; val zs = map Bound (l - 1 downto 0); in fold_rev (Term.abs o pair "z") Us (list_comb (p, make_bool_args' bs i @ make_args argTs ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us)))) end | NONE => (case t of t1 $ t2 => subst t1 $ subst t2 | Abs (x, T, u) => Abs (x, T, subst u) | _ => t)); (* transform an introduction rule into a conjunction *) (* [| p_i t; ... |] ==> p_j u *) (* is transformed into *) (* b_j & x_j = u & p b_j t & ... *) fun transform_rule r = let val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); val ps = make_bool_args HOLogic.mk_not I bs i @ map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @ map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r); in fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps) end; (* make a disjunction of all introduction rules *) val fp_fun = fold_rev lambda (p :: bs @ xs) (if null intr_ts then \<^term>\False\ else foldr1 HOLogic.mk_disj (map transform_rule intr_ts)); (* add definition of recursive predicates to theory *) val is_auxiliary = length cs > 1; val rec_binding = if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name; val rec_name = Binding.name_of rec_binding; val internals = Config.get lthy inductive_internals; val ((rec_const, (_, fp_def)), lthy') = lthy |> is_auxiliary ? Proof_Context.concealed |> Local_Theory.define ((rec_binding, case cnames_syn of [(_, mx)] => mx | _ => NoSyn), ((Thm.make_def_binding internals rec_binding, @{attributes [nitpick_unfold]}), fold_rev lambda params (Const (fp_name, (predT --> predT) --> predT) $ fp_fun))) ||> Proof_Context.restore_naming lthy; val fp_def' = Simplifier.rewrite (put_simpset HOL_basic_ss lthy' addsimps [fp_def]) (Thm.cterm_of lthy' (list_comb (rec_const, params))); val specs = if is_auxiliary then map_index (fn (i, ((b, mx), c)) => let val Ts = arg_types_of (length params) c; val xs = map Free (Variable.variant_frees lthy' intr_ts (mk_names "x" (length Ts) ~~ Ts)); in ((b, mx), ((Thm.make_def_binding internals b, []), fold_rev lambda (params @ xs) (list_comb (rec_const, params @ make_bool_args' bs i @ make_args argTs (xs ~~ Ts))))) end) (cnames_syn ~~ cs) else []; val (consts_defs, lthy'') = lthy' |> fold_map Local_Theory.define specs; val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs); val (_, ctxt'') = Variable.add_fixes (map (fst o dest_Free) params) lthy''; val mono = prove_mono quiet_mode skip_mono predT fp_fun monos ctxt''; val (_, lthy''') = lthy'' |> Local_Theory.note ((if internals then Binding.qualify true rec_name (Binding.name "mono") else Binding.empty, []), Proof_Context.export ctxt'' lthy'' [mono]); in (lthy''', Proof_Context.transfer (Proof_Context.theory_of lthy''') ctxt'', rec_binding, mono, fp_def', map (#2 o #2) consts_defs, list_comb (rec_const, params), preds, argTs, bs, xs) end; fun declare_rules rec_binding coind no_ind spec_name cnames preds intrs intr_bindings intr_atts elims eqs raw_induct lthy = let val rec_name = Binding.name_of rec_binding; fun rec_qualified qualified = Binding.qualify qualified rec_name; val intr_names = map Binding.name_of intr_bindings; val ind_case_names = if forall (equal "") intr_names then [] else [Attrib.case_names intr_names]; val induct = if coind then (raw_induct, [Attrib.case_names [rec_name], Attrib.case_conclusion (rec_name, intr_names), Attrib.consumes (1 - Thm.nprems_of raw_induct), - Attrib.internal (K (Induct.coinduct_pred (hd cnames)))]) + Attrib.internal \<^here> (K (Induct.coinduct_pred (hd cnames)))]) else if no_ind orelse length cnames > 1 then (raw_induct, ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))]) else (raw_induct RSN (2, rev_mp), ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))]); val (intrs', lthy1) = lthy |> Spec_Rules.add spec_name (if coind then Spec_Rules.Co_Inductive else Spec_Rules.Inductive) preds intrs |> Local_Theory.notes (map (rec_qualified false) intr_bindings ~~ intr_atts ~~ map (fn th => [([th], @{attributes [Pure.intro?]})]) intrs) |>> map (hd o snd); val (((_, elims'), (_, [induct'])), lthy2) = lthy1 |> Local_Theory.note ((rec_qualified true (Binding.name "intros"), []), intrs') ||>> fold_map (fn (name, (elim, cases, k)) => Local_Theory.note ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"), ((if forall (equal "") cases then [] else [Attrib.case_names cases]) @ [Attrib.consumes (1 - Thm.nprems_of elim), Attrib.constraints k, - Attrib.internal (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})), + Attrib.internal \<^here> (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})), [elim]) #> apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>> Local_Theory.note ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), #2 induct), [rulify lthy1 (#1 induct)]); val (eqs', lthy3) = lthy2 |> fold_map (fn (name, eq) => Local_Theory.note ((Binding.qualify true (Long_Name.base_name name) (Binding.name "simps"), - [Attrib.internal (K equation_add_permissive)]), [eq]) + [Attrib.internal \<^here> (K equation_add_permissive)]), [eq]) #> apfst (hd o snd)) (if null eqs then [] else (cnames ~~ eqs)) val (inducts, lthy4) = if no_ind orelse coind then ([], lthy3) else let val inducts = cnames ~~ Project_Rule.projects lthy3 (1 upto length cnames) induct' in lthy3 |> Local_Theory.notes [((rec_qualified true (Binding.name "inducts"), []), inducts |> map (fn (name, th) => ([th], ind_case_names @ [Attrib.consumes (1 - Thm.nprems_of th), - Attrib.internal (K (Induct.induct_pred name))])))] |>> snd o hd + Attrib.internal \<^here> (K (Induct.induct_pred name))])))] |>> snd o hd end; in (intrs', elims', eqs', induct', inducts, lthy4) end; type flags = {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool}; type add_ind_def = flags -> term list -> (Attrib.binding * term) list -> thm list -> term list -> (binding * mixfix) list -> local_theory -> result * local_theory; fun add_ind_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono} cs intros monos params cnames_syn lthy = let val _ = null cnames_syn andalso error "No inductive predicates given"; val names = map (Binding.name_of o fst) cnames_syn; val _ = message (quiet_mode andalso not verbose) ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names); val spec_name = Binding.conglomerate (map #1 cnames_syn); val cnames = map (Local_Theory.full_name lthy o #1) cnames_syn; (* FIXME *) val ((intr_names, intr_atts), intr_ts) = apfst split_list (split_list (map (check_rule lthy cs params) intros)); val (lthy1, lthy2, rec_binding, mono, fp_def, rec_preds_defs, rec_const, preds, argTs, bs, xs) = mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy; val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs) intr_ts rec_preds_defs lthy2 lthy1; val elims = if no_elim then [] else prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names) unfold rec_preds_defs lthy2 lthy1; val raw_induct = zero_var_indexes (if no_ind then Drule.asm_rl else if coind then prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono fp_def rec_preds_defs lthy2 lthy1 else prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def rec_preds_defs lthy2 lthy1); val eqs = if no_elim then [] else prove_eqs quiet_mode cs params intr_ts intrs elims lthy2 lthy1; val elims' = map (fn (th, ns, i) => (rulify lthy1 th, ns, i)) elims; val intrs' = map (rulify lthy1) intrs; val (intrs'', elims'', eqs', induct, inducts, lthy3) = declare_rules rec_binding coind no_ind spec_name cnames preds intrs' intr_names intr_atts elims' eqs raw_induct lthy1; val result = {preds = preds, intrs = intrs'', elims = elims'', raw_induct = rulify lthy3 raw_induct, induct = induct, inducts = inducts, eqs = eqs'}; val lthy4 = lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => let val result' = transform_result phi result; in put_inductives ({names = cnames, coind = coind}, result') end); in (result, lthy4) end; (* external interfaces *) fun gen_add_inductive mk_def flags cnames_syn pnames spec monos lthy = let (* abbrevs *) val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy; fun get_abbrev ((name, atts), t) = if can (Logic.strip_assums_concl #> Logic.dest_equals) t then let val _ = Binding.is_empty name andalso null atts orelse error "Abbreviations may not have names or attributes"; val ((x, T), rhs) = Local_Defs.abs_def (snd (Local_Defs.cert_def ctxt1 (K []) t)); val var = (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of NONE => error ("Undeclared head of abbreviation " ^ quote x) | SOME ((b, T'), mx) => if T <> T' then error ("Bad type specification for abbreviation " ^ quote x) else (b, mx)); in SOME (var, rhs) end else NONE; val abbrevs = map_filter get_abbrev spec; val bs = map (Binding.name_of o fst o fst) abbrevs; (* predicates *) val pre_intros = filter_out (is_some o get_abbrev) spec; val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn; val cs = map (Free o apfst Binding.name_of o fst) cnames_syn'; val ps = map Free pnames; val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn'); val ctxt3 = ctxt2 |> fold (snd oo Local_Defs.fixed_abbrev) abbrevs; val expand = Assumption.export_term ctxt3 lthy #> Proof_Context.cert_term lthy; fun close_rule r = fold (Logic.all o Free) (fold_aterms (fn t as Free (v as (s, _)) => if Variable.is_fixed ctxt1 s orelse member (op =) ps t then I else insert (op =) v | _ => I) r []) r; val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros; val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn'; in lthy |> mk_def flags cs intros monos ps preds ||> fold (snd oo Local_Theory.abbrev Syntax.mode_default) abbrevs end; fun gen_add_inductive_cmd mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos lthy = let val ((vars, intrs), _) = lthy |> Proof_Context.set_mode Proof_Context.mode_abbrev |> Specification.read_multi_specs (cnames_syn @ pnames_syn) intro_srcs; val (cs, ps) = chop (length cnames_syn) vars; val monos = Attrib.eval_thms lthy raw_monos; val flags = {quiet_mode = false, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false, skip_mono = false}; in lthy |> gen_add_inductive mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos end; val add_inductive = gen_add_inductive add_ind_def; val add_inductive_cmd = gen_add_inductive_cmd add_ind_def; (* read off arities of inductive predicates from raw induction rule *) fun arities_of induct = map (fn (_ $ t $ u) => (fst (dest_Const (head_of t)), length (snd (strip_comb u)))) (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct))); (* read off parameters of inductive predicate from raw induction rule *) fun params_of induct = let val (_ $ t $ u :: _) = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)); val (_, ts) = strip_comb t; val (_, us) = strip_comb u; in List.take (ts, length ts - length us) end; val pname_of_intr = Thm.concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst; (* partition introduction rules according to predicate name *) fun gen_partition_rules f induct intros = fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros (map (rpair [] o fst) (arities_of induct)); val partition_rules = gen_partition_rules I; fun partition_rules' induct = gen_partition_rules fst induct; fun unpartition_rules intros xs = fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r) (fn x :: xs => (x, xs)) #>> the) intros xs |> fst; (* infer order of variables in intro rules from order of quantifiers in elim rule *) fun infer_intro_vars thy elim arity intros = let val _ :: cases = Thm.prems_of elim; val used = map (fst o fst) (Term.add_vars (Thm.prop_of elim) []); fun mtch (t, u) = let val params = Logic.strip_params t; val vars = map (Var o apfst (rpair 0)) (Name.variant_list used (map fst params) ~~ map snd params); val ts = map (curry subst_bounds (rev vars)) (List.drop (Logic.strip_assums_hyp t, arity)); val us = Logic.strip_imp_prems u; val tab = fold (Pattern.first_order_match thy) (ts ~~ us) (Vartab.empty, Vartab.empty); in map (Envir.subst_term tab) vars end in map (mtch o apsnd Thm.prop_of) (cases ~~ intros) end; (** outer syntax **) fun gen_ind_decl mk_def coind = Parse.vars -- Parse.for_fixes -- Scan.optional Parse_Spec.where_multi_specs [] -- Scan.optional (\<^keyword>\monos\ |-- Parse.!!! Parse.thms1) [] >> (fn (((preds, params), specs), monos) => (snd o gen_add_inductive_cmd mk_def true coind preds params specs monos)); val ind_decl = gen_ind_decl add_ind_def; val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive\ "define inductive predicates" (ind_decl false); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinductive\ "define coinductive predicates" (ind_decl true); val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_cases\ "create simplified instances of elimination rules" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_cases_cmd)); val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_simps\ "create simplification rules for inductive predicates" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_simps_cmd)); val _ = Outer_Syntax.command \<^command_keyword>\print_inductives\ "print (co)inductive definitions and monotonicity rules" (Parse.opt_bang >> (fn b => Toplevel.keep (print_inductives b o Toplevel.context_of))); end; diff --git a/src/HOL/Tools/inductive_set.ML b/src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML +++ b/src/HOL/Tools/inductive_set.ML @@ -1,560 +1,560 @@ (* Title: HOL/Tools/inductive_set.ML Author: Stefan Berghofer, TU Muenchen Wrapper for defining inductive sets using package for inductive predicates, including infrastructure for converting between predicates and sets. *) signature INDUCTIVE_SET = sig val to_set_att: thm list -> attribute val to_pred_att: thm list -> attribute val to_pred : thm list -> Context.generic -> thm -> thm val pred_set_conv_att: attribute val add_inductive: Inductive.flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> Inductive.result * local_theory val add_inductive_cmd: bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> Inductive.result * local_theory val mono_add: attribute val mono_del: attribute end; structure Inductive_Set: INDUCTIVE_SET = struct (***********************************************************************************) (* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *) (* and (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y}) *) (* used for converting "strong" (co)induction rules *) (***********************************************************************************) val anyt = Free ("t", TFree ("'t", [])); fun strong_ind_simproc tab = Simplifier.make_simproc \<^context> "strong_ind" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let fun close p t f = let val vs = Term.add_vars t [] in Thm.instantiate' [] (rev (map (SOME o Thm.cterm_of ctxt o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; fun mkop \<^const_name>\HOL.conj\ T x = SOME (Const (\<^const_name>\Lattices.inf\, T --> T --> T), x) | mkop \<^const_name>\HOL.disj\ T x = SOME (Const (\<^const_name>\Lattices.sup\, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T in HOLogic.Collect_const U $ HOLogic.mk_ptupleabs (HOLogic.flat_tuple_paths p) U HOLogic.boolT t end; fun decomp (Const (s, _) $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) = mkop s T (m, p, S, mk_collect p T (head_of u)) | decomp (Const (s, _) $ u $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S)) = mkop s T (m, p, mk_collect p T (head_of u), S) | decomp _ = NONE; val simp = full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv}) 1; fun mk_rew t = (case strip_abs_vars t of [] => NONE | xs => (case decomp (strip_abs_body t) of NONE => NONE | SOME (bop, (m, p, S, S')) => SOME (close (Goal.prove ctxt [] []) (Logic.mk_equals (t, fold_rev Term.abs xs (m $ p $ (bop $ S $ S')))) (K (EVERY [resolve_tac ctxt [eq_reflection] 1, REPEAT (resolve_tac ctxt @{thms ext} 1), resolve_tac ctxt @{thms iffI} 1, EVERY [eresolve_tac ctxt @{thms conjE} 1, resolve_tac ctxt @{thms IntI} 1, simp, simp, eresolve_tac ctxt @{thms IntE} 1, resolve_tac ctxt @{thms conjI} 1, simp, simp] ORELSE EVERY [eresolve_tac ctxt @{thms disjE} 1, resolve_tac ctxt @{thms UnI1} 1, simp, resolve_tac ctxt @{thms UnI2} 1, simp, eresolve_tac ctxt @{thms UnE} 1, resolve_tac ctxt @{thms disjI1} 1, simp, resolve_tac ctxt @{thms disjI2} 1, simp]]))) handle ERROR _ => NONE)) in (case strip_comb (Thm.term_of ct) of (h as Const (name, _), ts) => if Symtab.defined tab name then let val rews = map mk_rew ts in if forall is_none rews then NONE else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1) (map2 (fn SOME r => K r | NONE => Thm.reflexive o Thm.cterm_of ctxt) rews ts) (Thm.reflexive (Thm.cterm_of ctxt h))) end else NONE | _ => NONE) end}; (* only eta contract terms occurring as arguments of functions satisfying p *) fun eta_contract p = let fun eta b (Abs (a, T, body)) = (case eta b body of body' as (f $ Bound 0) => if Term.is_dependent f orelse not b then Abs (a, T, body') else incr_boundvars ~1 f | body' => Abs (a, T, body')) | eta b (t $ u) = eta b t $ eta (p (head_of t)) u | eta b t = t in eta false end; fun eta_contract_thm ctxt p = Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct => Thm.transitive (Thm.eta_conversion ct) (Thm.symmetric (Thm.eta_conversion (Thm.cterm_of ctxt (eta_contract p (Thm.term_of ct))))))); (***********************************************************) (* rules for converting between predicate and set notation *) (* *) (* rules for converting predicates to sets have the form *) (* P (%x y. (x, y) : s) = (%x y. (x, y) : S s) *) (* *) (* rules for converting sets to predicates have the form *) (* S {(x, y). p x y} = {(x, y). P p x y} *) (* *) (* where s and p are parameters *) (***********************************************************) structure Data = Generic_Data ( type T = {(* rules for converting predicates to sets *) to_set_simps: thm list, (* rules for converting sets to predicates *) to_pred_simps: thm list, (* arities of functions of type t set => ... => u set *) set_arities: (typ * (int list list option list * int list list option)) list Symtab.table, (* arities of functions of type (t => ... => bool) => u => ... => bool *) pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table}; val empty = {to_set_simps = [], to_pred_simps = [], set_arities = Symtab.empty, pred_arities = Symtab.empty}; fun merge ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1, set_arities = set_arities1, pred_arities = pred_arities1}, {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2, set_arities = set_arities2, pred_arities = pred_arities2}) : T = {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2), to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2), set_arities = Symtab.merge_list (op =) (set_arities1, set_arities2), pred_arities = Symtab.merge_list (op =) (pred_arities1, pred_arities2)}; ); fun name_type_of (Free p) = SOME p | name_type_of (Const p) = SOME p | name_type_of _ = NONE; fun map_type f (Free (s, T)) = Free (s, f T) | map_type f (Var (ixn, T)) = Var (ixn, f T) | map_type f _ = error "map_type"; fun find_most_specific is_inst f eq xs T = find_first (fn U => is_inst (T, f U) andalso forall (fn U' => eq (f U, f U') orelse not (is_inst (T, f U') andalso is_inst (f U', f U))) xs) xs; fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of NONE => NONE | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T; fun lookup_rule thy f rules = find_most_specific (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules; fun infer_arities thy arities (optf, t) fs = case strip_comb t of (Abs (_, _, u), []) => infer_arities thy arities (NONE, u) fs | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of SOME (SOME (_, (arity, _))) => (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs handle General.Subscript => error "infer_arities: bad term") | _ => fold (infer_arities thy arities) (map (pair NONE) ts) (case optf of NONE => fs | SOME f => AList.update op = (u, the_default f (Option.map (fn g => inter (op =) g f) (AList.lookup op = fs u))) fs)); (**************************************************************) (* derive the to_pred equation from the to_set equation *) (* *) (* 1. instantiate each set parameter with {(x, y). p x y} *) (* 2. apply %P. {(x, y). P x y} to both sides of the equation *) (* 3. simplify *) (**************************************************************) fun mk_to_pred_inst ctxt fs = map (fn (x, ps) => let val (Ts, T) = strip_type (fastype_of x); val U = HOLogic.dest_setT T; val x' = map_type (K (Ts @ HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x; in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs ps U HOLogic.boolT (list_comb (x', map Bound (length Ts - 1 downto 0)))))) end) fs; fun mk_to_pred_eq ctxt p fs optfs' T thm = let val insts = mk_to_pred_inst ctxt fs; val thm' = Thm.instantiate (TVars.empty, Vars.make insts) thm; val thm'' = (case optfs' of NONE => thm' RS sym | SOME fs' => let val U = HOLogic.dest_setT (body_type T); val Ts = HOLogic.strip_ptupleT fs' U; val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong; val (Var (arg_cong_f, _), _) = arg_cong' |> Thm.concl_of |> dest_comb |> snd |> strip_comb |> snd |> hd |> dest_comb; in thm' RS (infer_instantiate ctxt [(arg_cong_f, Thm.cterm_of ctxt (Abs ("P", Ts ---> HOLogic.boolT, HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs' U HOLogic.boolT (Bound 0))))] arg_cong' RS sym) end) in Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv} addsimprocs [\<^simproc>\Collect_mem\]) thm'' |> zero_var_indexes |> eta_contract_thm ctxt (equal p) end; (**** declare rules for converting predicates to sets ****) exception Malformed of string; fun add context thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) = (case Thm.prop_of thm of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ lhs $ rhs) => (case body_type T of \<^typ>\bool\ => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; fun factors_of t fs = case strip_abs_body t of Const (\<^const_name>\Set.member\, _) $ u $ S => if is_Free S orelse is_Var S then let val ps = HOLogic.flat_tuple_paths u in (SOME ps, (S, ps) :: fs) end else (NONE, fs) | _ => (NONE, fs); val (h, ts) = strip_comb lhs val (pfs, fs) = fold_map factors_of ts []; val ((h', ts'), fs') = (case rhs of Abs _ => (case strip_abs_body rhs of Const (\<^const_name>\Set.member\, _) $ u $ S => (strip_comb S, SOME (HOLogic.flat_tuple_paths u)) | _ => raise Malformed "member symbol on right-hand side expected") | _ => (strip_comb rhs, NONE)) in case (name_type_of h, name_type_of h') of (SOME (s, T), SOME (s', T')) => if exists (fn (U, _) => Sign.typ_instance thy (T', U) andalso Sign.typ_instance thy (U, T')) (Symtab.lookup_list set_arities s') then (if Context_Position.is_really_visible ctxt then warning ("Ignoring conversion rule for operator " ^ s') else (); tab) else {to_set_simps = Thm.trim_context thm :: to_set_simps, to_pred_simps = Thm.trim_context (mk_to_pred_eq ctxt h fs fs' T' thm) :: to_pred_simps, set_arities = Symtab.insert_list op = (s', (T', (map (AList.lookup op = fs) ts', fs'))) set_arities, pred_arities = Symtab.insert_list op = (s, (T, (pfs, fs'))) pred_arities} | _ => raise Malformed "set / predicate constant expected" end | _ => raise Malformed "equation between predicates expected") | _ => raise Malformed "equation expected") handle Malformed msg => let val ctxt = Context.proof_of context val _ = if Context_Position.is_really_visible ctxt then warning ("Ignoring malformed set / predicate conversion rule: " ^ msg ^ "\n" ^ Thm.string_of_thm ctxt thm) else (); in tab end; val pred_set_conv_att = Thm.declaration_attribute (fn thm => fn ctxt => Data.map (add ctxt thm) ctxt); (**** convert theorem in set notation to predicate notation ****) fun is_pred tab t = case Option.map (Symtab.lookup tab o fst) (name_type_of t) of SOME (SOME _) => true | _ => false; fun to_pred_simproc rules = let val rules' = map mk_meta_eq rules in Simplifier.make_simproc \<^context> "to_pred" {lhss = [anyt], proc = fn _ => fn ctxt => fn ct => lookup_rule (Proof_Context.theory_of ctxt) (Thm.prop_of #> Logic.dest_equals) rules' (Thm.term_of ct)} end; fun to_pred_proc thy rules t = case lookup_rule thy I rules t of NONE => NONE | SOME (lhs, rhs) => SOME (Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs); fun to_pred thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_pred_simps, set_arities, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy set_arities (NONE, Thm.prop_of thm) []); (* instantiate each set parameter with {(x, y). p x y} *) val insts = mk_to_pred_inst ctxt fs in thm |> Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimprocs [to_pred_simproc (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: map (Thm.transfer thy) to_pred_simps)]) |> eta_contract_thm ctxt (is_pred pred_arities) |> Rule_Cases.save thm end; val to_pred_att = Thm.rule_attribute [] o to_pred; (**** convert theorem in predicate notation to set notation ****) fun to_set thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_set_simps, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy pred_arities (NONE, Thm.prop_of thm) []); (* instantiate each predicate parameter with %x y. (x, y) : s *) val insts = map (fn (x, ps) => let val Ts = binder_types (fastype_of x); val l = length Ts; val k = length ps; val (Rs, Us) = chop (l - k - 1) Ts; val T = HOLogic.mk_ptupleT ps Us; val x' = map_type (K (Rs ---> HOLogic.mk_setT T)) x in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple ps T (map Bound (k downto 0)), list_comb (x', map Bound (l - 1 downto k + 1)))))) end) fs; in thm |> Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps to_set_simps addsimprocs [strong_ind_simproc pred_arities, \<^simproc>\Collect_mem\]) |> Rule_Cases.save thm end; val to_set_att = Thm.rule_attribute [] o to_set; (**** definition of inductive sets ****) fun add_ind_set_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono} cs intros monos params cnames_syn lthy = let val thy = Proof_Context.theory_of lthy; val {set_arities, pred_arities, to_pred_simps, ...} = Data.get (Context.Proof lthy); fun infer (Abs (_, _, t)) = infer t | infer (Const (\<^const_name>\Set.member\, _) $ t $ u) = infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u) | infer (t $ u) = infer t #> infer u | infer _ = I; val new_arities = filter_out (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 0 | _ => false) (fold (snd #> infer) intros []); val params' = map (fn x => (case AList.lookup op = new_arities x of SOME fs => let val T = HOLogic.dest_setT (fastype_of x); val Ts = HOLogic.strip_ptupleT fs T; val x' = map_type (K (Ts ---> HOLogic.boolT)) x in (x, (x', (HOLogic.Collect_const T $ HOLogic.mk_ptupleabs fs T HOLogic.boolT x', fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs T (map Bound (length fs downto 0)), x))))) end | NONE => (x, (x, (x, x))))) params; val (params1, (params2, params3)) = params' |> map snd |> split_list ||> split_list; val paramTs = map fastype_of params; (* equations for converting sets to predicates *) val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) => let val fs = the_default [] (AList.lookup op = new_arities c); val (Us, U) = strip_type T |> apsnd HOLogic.dest_setT; val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks [Pretty.str "Argument types", Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)), Pretty.str ("of " ^ s ^ " do not agree with types"), Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)), Pretty.str "of declared parameters"])); val Ts = HOLogic.strip_ptupleT fs U; val c' = Free (s ^ "p", map fastype_of params1 @ Ts ---> HOLogic.boolT) in ((c', (fs, U, Ts)), (list_comb (c, params2), HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (c', params1)))) end) |> split_list |>> split_list; val eqns' = eqns @ map (Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: to_pred_simps); (* predicate version of the introduction rules *) val intros' = map (fn (name_atts, t) => (name_atts, t |> map_aterms (fn u => (case AList.lookup op = params' u of SOME (_, (u', _)) => u' | NONE => u)) |> Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |> eta_contract (member op = cs' orf is_pred pred_arities))) intros; val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn; val monos' = map (to_pred [] (Context.Proof lthy)) monos; val ({preds, intrs, elims, raw_induct, eqs, ...}, lthy1) = Inductive.add_ind_def {quiet_mode = quiet_mode, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono} cs' intros' monos' params1 cnames_syn' lthy; (* define inductive sets using previously defined predicates *) val (defs, lthy2) = lthy1 |> fold_map Local_Theory.define (map (fn (((b, mx), (fs, U, _)), p) => ((b, mx), ((Thm.def_binding b, []), fold_rev lambda params (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (p, params3)))))) (cnames_syn ~~ cs_info ~~ preds)); (* prove theorems for converting predicate to set notation *) val lthy3 = fold (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn lthy => let val conv_thm = Goal.prove lthy (map (fst o dest_Free) params) [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (p, params3), fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs U (map Bound (length fs downto 0)), list_comb (c, params)))))) (K (REPEAT (resolve_tac lthy @{thms ext} 1) THEN simp_tac (put_simpset HOL_basic_ss lthy addsimps [def, @{thm mem_Collect_eq}, @{thm case_prod_conv}]) 1)) in lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"), - [Attrib.internal (K pred_set_conv_att)]), + [Attrib.internal \<^here> (K pred_set_conv_att)]), [conv_thm]) |> snd end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2; (* convert theorems to set notation *) val rec_name = if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name; val cnames = map (Local_Theory.full_name lthy3 o #1) cnames_syn; (* FIXME *) val spec_name = Binding.conglomerate (map #1 cnames_syn); val (intr_names, intr_atts) = split_list (map fst intros); val raw_induct' = to_set [] (Context.Proof lthy3) raw_induct; val (intrs', elims', eqs', induct, inducts, lthy4) = Inductive.declare_rules rec_name coind no_ind spec_name cnames (map fst defs) (map (to_set [] (Context.Proof lthy3)) intrs) intr_names intr_atts (map (fn th => (to_set [] (Context.Proof lthy3) th, map (fst o fst) (fst (Rule_Cases.get th)), Rule_Cases.get_constraints th)) elims) (map (to_set [] (Context.Proof lthy3)) eqs) raw_induct' lthy3; in ({intrs = intrs', elims = elims', induct = induct, inducts = inducts, raw_induct = raw_induct', preds = map fst defs, eqs = eqs'}, lthy4) end; val add_inductive = Inductive.gen_add_inductive add_ind_set_def; val add_inductive_cmd = Inductive.gen_add_inductive_cmd add_ind_set_def; fun mono_att att = Thm.declaration_attribute (fn thm => fn context => Thm.attribute_declaration att (to_pred [] context thm) context); val mono_add = mono_att Inductive.mono_add; val mono_del = mono_att Inductive.mono_del; (** package setup **) (* attributes *) val _ = Theory.setup (Attrib.setup \<^binding>\pred_set_conv\ (Scan.succeed pred_set_conv_att) "declare rules for converting between predicate and set notation" #> Attrib.setup \<^binding>\to_set\ (Attrib.thms >> to_set_att) "convert rule to set notation" #> Attrib.setup \<^binding>\to_pred\ (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #> Attrib.setup \<^binding>\mono_set\ (Attrib.add_del mono_add mono_del) "declare of monotonicity rule for set operators"); (* commands *) val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def; val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_set\ "define inductive sets" (ind_set_decl false); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinductive_set\ "define coinductive sets" (ind_set_decl true); end; diff --git a/src/HOL/Tools/semiring_normalizer.ML b/src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML +++ b/src/HOL/Tools/semiring_normalizer.ML @@ -1,885 +1,885 @@ (* Title: HOL/Tools/semiring_normalizer.ML Author: Amine Chaieb, TU Muenchen Normalization of expressions in semirings. *) signature SEMIRING_NORMALIZER = sig type entry val match: Proof.context -> cterm -> entry option val the_semiring: Proof.context -> thm -> cterm list * thm list val the_ring: Proof.context -> thm -> cterm list * thm list val the_field: Proof.context -> thm -> cterm list * thm list val the_idom: Proof.context -> thm -> thm list val the_ideal: Proof.context -> thm -> thm list val declare: thm -> {semiring: term list * thm list, ring: term list * thm list, field: term list * thm list, idom: thm list, ideal: thm list} -> local_theory -> local_theory val semiring_normalize_conv: Proof.context -> conv val semiring_normalize_ord_conv: Proof.context -> cterm ord -> conv val semiring_normalize_wrapper: Proof.context -> entry -> conv val semiring_normalize_ord_wrapper: Proof.context -> entry -> cterm ord -> conv val semiring_normalizers_conv: cterm list -> cterm list * thm list -> cterm list * thm list -> cterm list * thm list -> (cterm -> bool) * conv * conv * conv -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} val semiring_normalizers_ord_wrapper: Proof.context -> entry -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} end structure Semiring_Normalizer: SEMIRING_NORMALIZER = struct (** data **) type entry = {vars: cterm list, semiring: cterm list * thm list, ring: cterm list * thm list, field: cterm list * thm list, idom: thm list, ideal: thm list} * {is_const: cterm -> bool, dest_const: cterm -> Rat.rat, mk_const: ctyp -> Rat.rat -> cterm, conv: Proof.context -> cterm -> thm}; structure Data = Generic_Data ( type T = (thm * entry) list; val empty = []; fun merge data = AList.merge Thm.eq_thm (K true) data; ); fun the_rules ctxt = fst o the o AList.lookup Thm.eq_thm (Data.get (Context.Proof ctxt)) val the_semiring = #semiring oo the_rules val the_ring = #ring oo the_rules val the_field = #field oo the_rules val the_idom = #idom oo the_rules val the_ideal = #ideal oo the_rules fun match ctxt tm = let fun match_inst ({vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal}, fns) pat = let fun h instT = let val substT = Thm.instantiate (instT, Vars.empty); val substT_cterm = Drule.cterm_rule substT; val vars' = map substT_cterm vars; val semiring' = (map substT_cterm sr_ops, map substT sr_rules); val ring' = (map substT_cterm r_ops, map substT r_rules); val field' = (map substT_cterm f_ops, map substT f_rules); val idom' = map substT idom; val ideal' = map substT ideal; val result = ({vars = vars', semiring = semiring', ring = ring', field = field', idom = idom', ideal = ideal'}, fns); in SOME result end in (case try Thm.match (pat, tm) of NONE => NONE | SOME (instT, _) => h instT) end; fun match_struct (_, entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) = get_first (match_inst entry) (sr_ops @ r_ops @ f_ops); in get_first match_struct (Data.get (Context.Proof ctxt)) end; (* extra-logical functions *) val semiring_norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms semiring_norm}); val semiring_funs = {is_const = can HOLogic.dest_number o Thm.term_of, dest_const = (fn ct => Rat.of_int (snd (HOLogic.dest_number (Thm.term_of ct) handle TERM _ => error "ring_dest_const"))), mk_const = (fn cT => fn x => Numeral.mk_cnumber cT (case Rat.dest x of (i, 1) => i | _ => error "int_of_rat: bad int")), conv = (fn ctxt => Simplifier.rewrite (put_simpset semiring_norm_ss ctxt) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms numeral_One}))}; val divide_const = Thm.cterm_of \<^context> (Logic.varify_global \<^term>\(/)\); val [divide_tvar] = Term.add_tvars (Thm.term_of divide_const) []; val field_funs = let fun numeral_is_const ct = case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b => can HOLogic.dest_number a andalso can HOLogic.dest_number b | Const (\<^const_name>\Fields.inverse\,_)$t => can HOLogic.dest_number t | t => can HOLogic.dest_number t fun dest_const ct = ((case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b=> Rat.make (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | Const (\<^const_name>\Fields.inverse\,_)$t => Rat.inv (Rat.of_int (snd (HOLogic.dest_number t))) | t => Rat.of_int (snd (HOLogic.dest_number t))) handle TERM _ => error "ring_dest_const") fun mk_const cT x = let val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber cT a else Thm.apply (Thm.apply (Thm.instantiate_cterm (TVars.make1 (divide_tvar, cT), Vars.empty) divide_const) (Numeral.mk_cnumber cT a)) (Numeral.mk_cnumber cT b) end in {is_const = numeral_is_const, dest_const = dest_const, mk_const = mk_const, conv = Numeral_Simprocs.field_comp_conv} end; (* logical content *) val semiringN = "semiring"; val ringN = "ring"; val fieldN = "field"; val idomN = "idom"; fun declare raw_key {semiring = raw_semiring0, ring = raw_ring0, field = raw_field0, idom = raw_idom, ideal = raw_ideal} lthy = let val ctxt' = fold Proof_Context.augment (fst raw_semiring0 @ fst raw_ring0 @ fst raw_field0) lthy; val prepare_ops = apfst (Variable.export_terms ctxt' lthy #> map (Thm.cterm_of lthy)); val raw_semiring = prepare_ops raw_semiring0; val raw_ring = prepare_ops raw_ring0; val raw_field = prepare_ops raw_field0; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => fn context => let val ctxt = Context.proof_of context; val key = Morphism.thm phi raw_key; fun transform_ops_rules (ops, rules) = (map (Morphism.cterm phi) ops, Morphism.fact phi rules); val (sr_ops, sr_rules) = transform_ops_rules raw_semiring; val (r_ops, r_rules) = transform_ops_rules raw_ring; val (f_ops, f_rules) = transform_ops_rules raw_field; val idom = Morphism.fact phi raw_idom; val ideal = Morphism.fact phi raw_ideal; fun check kind name xs n = null xs orelse length xs = n orelse error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name); val check_ops = check "operations"; val check_rules = check "rules"; val _ = check_ops semiringN sr_ops 5 andalso check_rules semiringN sr_rules 36 andalso check_ops ringN r_ops 2 andalso check_rules ringN r_rules 2 andalso check_ops fieldN f_ops 2 andalso check_rules fieldN f_rules 2 andalso check_rules idomN idom 2; val mk_meta = Local_Defs.meta_rewrite_rule ctxt; val sr_rules' = map mk_meta sr_rules; val r_rules' = map mk_meta r_rules; val f_rules' = map mk_meta f_rules; fun rule i = nth sr_rules' (i - 1); val (cx, cy) = Thm.dest_binop (hd sr_ops); val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val ((clx, crx), (cly, cry)) = rule 13 |> Thm.rhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val ((ca, cb), (cc, cd)) = rule 20 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg; val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_arg; val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry]; val semiring = (sr_ops, sr_rules'); val ring = (r_ops, r_rules'); val field = (f_ops, f_rules'); val ideal' = map (Thm.symmetric o mk_meta) ideal in context |> Data.map (AList.update Thm.eq_thm (key, ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal'}, (if null f_ops then semiring_funs else field_funs)))) end) end; (** auxiliary **) fun is_comb ct = (case Thm.term_of ct of _ $ _ => true | _ => false); val concl = Thm.cprop_of #> Thm.dest_arg; fun is_binop ct ct' = (case Thm.term_of ct' of c $ _ $ _ => Thm.term_of ct aconv c | _ => false); fun dest_binop ct ct' = if is_binop ct ct' then Thm.dest_binop ct' else raise CTERM ("dest_binop: bad binop", [ct, ct']) fun inst_thm inst = Thm.instantiate (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) inst)); val dest_number = Thm.term_of #> HOLogic.dest_number #> snd; val is_number = can dest_number; fun numeral01_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One}]); fun zero1_numeral_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One} RS sym]); fun zerone_conv ctxt cv = zero1_numeral_conv ctxt then_conv cv then_conv numeral01_conv ctxt; val nat_add_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms arith_simps} @ @{thms diff_nat_numeral} @ @{thms rel_simps} @ @{thms if_False if_True Nat.add_0 add_Suc add_numeral_left Suc_eq_plus1} @ map (fn th => th RS sym) @{thms numerals}); fun nat_add_conv ctxt = zerone_conv ctxt (Simplifier.rewrite (put_simpset nat_add_ss ctxt)); val zeron_tm = \<^cterm>\0::nat\; val onen_tm = \<^cterm>\1::nat\; val true_tm = \<^cterm>\True\; (** normalizing conversions **) (* core conversion *) fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) (f_ops, f_rules) (is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) = let val [pthm_02, pthm_03, pthm_04, pthm_05, pthm_07, pthm_08, pthm_09, pthm_10, pthm_11, pthm_12, pthm_13, pthm_14, pthm_15, pthm_16, pthm_17, pthm_18, pthm_19, pthm_21, pthm_22, pthm_23, pthm_24, pthm_25, pthm_26, pthm_27, pthm_28, pthm_29, pthm_30, pthm_31, pthm_32, pthm_33, pthm_34, pthm_35, pthm_36, pthm_37, pthm_38, _] = sr_rules; val [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry] = vars; val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; val [add_tm, mul_tm, pow_tm] = map (Thm.dest_fun o Thm.dest_fun) [add_pat, mul_pat, pow_pat]; val dest_add = dest_binop add_tm val dest_mul = dest_binop mul_tm fun dest_pow tm = let val (l,r) = dest_binop pow_tm tm in if is_number r then (l,r) else raise CTERM ("dest_pow",[tm]) end; val is_add = is_binop add_tm val is_mul = is_binop mul_tm val (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, cx', cy') = (case (r_ops, r_rules) of ([sub_pat, neg_pat], [neg_mul, sub_add]) => let val sub_tm = Thm.dest_fun (Thm.dest_fun sub_pat) val neg_tm = Thm.dest_fun neg_pat val dest_sub = dest_binop sub_tm in (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, neg_mul |> concl |> Thm.dest_arg, sub_add |> concl |> Thm.dest_arg |> Thm.dest_arg) end | _ => (TrueI, TrueI, true_tm, true_tm, (fn t => (t,t)), true_tm, true_tm)); val (divide_inverse, divide_tm, inverse_tm) = (case (f_ops, f_rules) of ([divide_pat, inverse_pat], [div_inv, _]) => let val div_tm = funpow 2 Thm.dest_fun divide_pat val inv_tm = Thm.dest_fun inverse_pat in (div_inv, div_tm, inv_tm) end | _ => (TrueI, true_tm, true_tm)); in fn variable_ord => let (* Conversion for "x^n * x^m", with either x^n = x and/or x^m = x possible. *) (* Also deals with "const * const", but both terms must involve powers of *) (* the same variable, or both be constants, or behaviour may be incorrect. *) fun powvar_mul_conv ctxt tm = let val (l,r) = dest_mul tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_mul_conv tm else ((let val (lx,ln) = dest_pow l in ((let val (_, rn) = dest_pow r val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => (let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end)) end) handle CTERM _ => ((let val (rx,rn) = dest_pow r val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => inst_thm [(cx,l)] pthm_32 )) end; (* Remove "1 * m" from a monomial, and just leave m. *) fun monomial_deone th = (let val (l,r) = dest_mul(concl th) in if l aconvc one_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_13) else th end) handle CTERM _ => th; (* Conversion for "(monomial)^n", where n is a numeral. *) fun monomial_pow_conv ctxt = let fun monomial_pow tm bod ntm = if not(is_comb bod) then Thm.reflexive tm else if is_semiring_constant bod then semiring_pow_conv tm else let val (lopr,r) = Thm.dest_comb bod in if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 val (l,r) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv ctxt r)) end else if opr aconvc mul_tm then let val th1 = inst_thm [(cx,l),(cy,r),(cq,ntm)] pthm_33 val (xy,z) = Thm.dest_comb(concl th1) val (x,y) = Thm.dest_comb xy val thl = monomial_pow y l ntm val thr = monomial_pow z r ntm in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule x thl) thr) end else Thm.reflexive tm end end in fn tm => let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if not (opr aconvc pow_tm) orelse not(is_number r) then raise CTERM ("monomial_pow_conv", [tm]) else if r aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if r aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else monomial_deone(monomial_pow tm l r) end end; (* Multiplication of canonical monomials. *) fun monomial_mul_conv ctxt = let fun powvar tm = if is_semiring_constant tm then one_tm else ((let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then l else raise CTERM ("monomial_mul_conv",[tm]) end) handle CTERM _ => tm) (* FIXME !? *) fun vorder x y = if x aconvc y then 0 else if x aconvc one_tm then ~1 else if y aconvc one_tm then 1 else if is_less (variable_ord (x, y)) then ~1 else 1 fun monomial_mul tm l r = ((let val (lx,ly) = dest_mul l val vl = powvar lx in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 val th3 = Thm.transitive th1 th2 val (tm5,tm6) = Thm.dest_comb(concl th3) val (tm7,tm8) = Thm.dest_comb tm6 val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8 in Thm.transitive th3 (Drule.arg_cong_rule tm5 th4) end else let val th0 = if ord < 0 then pthm_16 else pthm_17 val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 in Thm.transitive th1 th2 end else if ord < 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else inst_thm [(ca,l),(cb,r)] pthm_09 end)) end) handle CTERM _ => (let val vl = powvar l in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2) end else if ord > 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else Thm.reflexive tm end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then powvar_mul_conv ctxt tm else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 else Thm.reflexive tm end)) end)) in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r) end end; (* Multiplication by monomial of a polynomial. *) fun polynomial_monomial_mul_conv ctxt = let fun pmm_conv tm = let val (l,r) = dest_mul tm in ((let val (y,z) = dest_add r val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv ctxt tm4)) (pmm_conv tm2) in Thm.transitive th1 th2 end) handle CTERM _ => monomial_mul_conv ctxt tm) end in pmm_conv end; (* Addition of two monomials identical except for constant multiples. *) fun monomial_add_conv tm = let val (l,r) = dest_add tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_add_conv tm else let val th1 = if is_mul l andalso is_semiring_constant(Thm.dest_arg1 l) then if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(ca,Thm.dest_arg1 l),(cm,Thm.dest_arg r), (cb,Thm.dest_arg1 r)] pthm_02 else inst_thm [(ca,Thm.dest_arg1 l),(cm,r)] pthm_03 else if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(cm,l),(ca,Thm.dest_arg1 r)] pthm_04 else inst_thm [(cm,r)] pthm_05 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4) val th3 = Thm.transitive th1 (Drule.fun_cong_rule th2 tm2) val tm5 = concl th3 in if (Thm.dest_arg1 tm5) aconvc zero_tm then Thm.transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11) else monomial_deone th3 end end; (* Ordering on monomials. *) fun striplist dest = let fun strip x acc = ((let val (l,r) = dest x in strip l (strip r acc) end) handle CTERM _ => x::acc) (* FIXME !? *) in fn x => strip x [] end; fun powervars tm = let val ptms = striplist dest_mul tm in if is_semiring_constant (hd ptms) then tl ptms else ptms end; val num_0 = 0; val num_1 = 1; fun dest_varpow tm = ((let val (x,n) = dest_pow tm in (x,dest_number n) end) handle CTERM _ => (tm,(if is_semiring_constant tm then num_0 else num_1))); val morder = let fun lexorder ls = case ls of ([],[]) => 0 | (_ ,[]) => ~1 | ([], _) => 1 | (((x1,n1)::vs1),((x2,n2)::vs2)) => (case variable_ord (x1, x2) of LESS => 1 | GREATER => ~1 | EQUAL => if n1 < n2 then ~1 else if n2 < n1 then 1 else lexorder (vs1, vs2)) in fn tm1 => fn tm2 => let val vdegs1 = map dest_varpow (powervars tm1) val vdegs2 = map dest_varpow (powervars tm2) val deg1 = fold (Integer.add o snd) vdegs1 num_0 val deg2 = fold (Integer.add o snd) vdegs2 num_0 in if deg1 < deg2 then ~1 else if deg1 > deg2 then 1 else lexorder (vdegs1, vdegs2) end end; (* Addition of two polynomials. *) fun polynomial_add_conv ctxt = let fun dezero_rule th = let val tm = concl th in if not(is_add tm) then th else let val (lopr,r) = Thm.dest_comb tm val l = Thm.dest_arg lopr in if l aconvc zero_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_07) else if r aconvc zero_tm then Thm.transitive th (inst_thm [(ca,l)] pthm_08) else th end end fun padd tm = let val (l,r) = dest_add tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_07 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_08 else if is_add l then let val (a,b) = dest_add l in if is_add r then let val (c,d) = dest_add r val ord = morder a c in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,c),(cd,d)] pthm_23 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4) in dezero_rule (Thm.transitive th1 (Thm.combination th2 (padd tm2))) end else (* ord <> 0*) let val th1 = if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else (* not (is_add r)*) let val ord = morder a r in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_26 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else (* ord <> 0*) if ord > 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end else (* not (is_add l)*) if is_add r then let val (c,d) = dest_add r val ord = morder l c in if ord = 0 then let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_28 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else if ord > 0 then Thm.reflexive tm else let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else let val ord = morder l r in if ord = 0 then monomial_add_conv tm else if ord > 0 then dezero_rule(Thm.reflexive tm) else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end in padd end; (* Multiplication of two polynomials. *) fun polynomial_mul_conv ctxt = let fun pmul tm = let val (l,r) = dest_mul tm in if not(is_add l) then polynomial_monomial_mul_conv ctxt tm else if not(is_add r) then let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 in Thm.transitive th1 (polynomial_monomial_mul_conv ctxt (concl th1)) end else let val (a,b) = dest_add l val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv ctxt tm4) val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2)) in Thm.transitive th3 (polynomial_add_conv ctxt (concl th3)) end end in fn tm => let val (l,r) = dest_mul tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_11 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_12 else if l aconvc one_tm then inst_thm [(ca,r)] pthm_13 else if r aconvc one_tm then inst_thm [(ca,l)] pthm_14 else pmul tm end end; (* Power of polynomial (optimized for the monomial and trivial cases). *) fun num_conv ctxt n = nat_add_conv ctxt (Thm.apply \<^cterm>\Suc\ (Numeral.mk_cnumber \<^ctyp>\nat\ (dest_number n - 1))) |> Thm.symmetric; fun polynomial_pow_conv ctxt = let fun ppow tm = let val (l,n) = dest_pow tm in if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else let val th1 = num_conv ctxt n val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 val (tm1,tm2) = Thm.dest_comb(concl th2) val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 in Thm.transitive th4 (polynomial_mul_conv ctxt (concl th4)) end end in fn tm => if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv ctxt tm end; (* Negation. *) fun polynomial_neg_conv ctxt tm = let val (l,r) = Thm.dest_comb tm in if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else let val th1 = inst_thm [(cx', r)] neg_mul val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1)) in Thm.transitive th2 (polynomial_monomial_mul_conv ctxt (concl th2)) end end; (* Subtraction. *) fun polynomial_sub_conv ctxt tm = let val (l,r) = dest_sub tm val th1 = inst_thm [(cx', l), (cy', r)] sub_add val (tm1,tm2) = Thm.dest_comb(concl th1) val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv ctxt tm2) in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv ctxt (concl th2))) end; (* Conversion from HOL term. *) fun polynomial_conv ctxt tm = if is_semiring_constant tm then semiring_add_conv tm else if not(is_comb tm) then Thm.reflexive tm else let val (lopr,r) = Thm.dest_comb tm in if lopr aconvc neg_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (polynomial_neg_conv ctxt (concl th1)) end else if lopr aconvc inverse_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (semiring_mul_conv (concl th1)) end else if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) r in Thm.transitive th1 (polynomial_pow_conv ctxt (concl th1)) end else if opr aconvc divide_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv ctxt) (Thm.rhs_of th1) in Thm.transitive th1 th2 end else if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val f = if opr aconvc add_tm then polynomial_add_conv ctxt else if opr aconvc mul_tm then polynomial_mul_conv ctxt else polynomial_sub_conv ctxt in Thm.transitive th1 (f (concl th1)) end else Thm.reflexive tm end end; in {main = polynomial_conv, add = polynomial_add_conv, mul = polynomial_mul_conv, pow = polynomial_pow_conv, neg = polynomial_neg_conv, sub = polynomial_sub_conv} end end; val nat_exp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps}) addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]); (* various normalizing conversions *) fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = let val pow_conv = Conv.arg_conv (Simplifier.rewrite (put_simpset nat_exp_ss ctxt)) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) then_conv conv ctxt val dat = (is_const, conv ctxt, conv ctxt, pow_conv) in semiring_normalizers_conv vars semiring ring field dat term_ord end; fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = #main (semiring_normalizers_ord_wrapper ctxt ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal}, {conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) term_ord) ctxt; fun semiring_normalize_wrapper ctxt data = semiring_normalize_ord_wrapper ctxt data Thm.term_ord; fun semiring_normalize_ord_conv ctxt ord tm = (case match ctxt tm of NONE => Thm.reflexive tm | SOME res => semiring_normalize_ord_wrapper ctxt res ord tm); fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt Thm.term_ord; end; diff --git a/src/HOL/Tools/typedef.ML b/src/HOL/Tools/typedef.ML --- a/src/HOL/Tools/typedef.ML +++ b/src/HOL/Tools/typedef.ML @@ -1,373 +1,373 @@ (* Title: HOL/Tools/typedef.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen Gordon/HOL-style type definitions: create a new syntactic type represented by a non-empty set. *) signature TYPEDEF = sig type info = {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} * {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, Rep_induct: thm, Abs_induct: thm} val transform_info: morphism -> info -> info val get_info: Proof.context -> string -> info list val get_info_global: theory -> string -> info list val interpretation: (string -> local_theory -> local_theory) -> theory -> theory type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding} val default_bindings: binding -> bindings val make_bindings: binding -> bindings option -> bindings val make_morphisms: binding -> (binding * binding) option -> bindings val overloaded: bool Config.T val add_typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> (Proof.context -> tactic) -> local_theory -> (string * info) * local_theory val add_typedef_global: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> (Proof.context -> tactic) -> theory -> (string * info) * theory val typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> local_theory -> Proof.state val typedef_cmd: {overloaded: bool} -> binding * (string * string option) list * mixfix -> string -> bindings option -> local_theory -> Proof.state end; structure Typedef: TYPEDEF = struct (** type definitions **) (* theory data *) type info = (*global part*) {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} * (*local part*) {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, Rep_induct: thm, Abs_induct: thm}; fun transform_info phi (info: info) = let val thm = Morphism.thm phi; val (global_info, {inhabited, type_definition, Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject, Rep_cases, Abs_cases, Rep_induct, Abs_induct}) = info; in (global_info, {inhabited = thm inhabited, type_definition = thm type_definition, Rep = thm Rep, Rep_inverse = thm Rep_inverse, Abs_inverse = thm Abs_inverse, Rep_inject = thm Rep_inject, Abs_inject = thm Abs_inject, Rep_cases = thm Rep_cases, Abs_cases = thm Abs_cases, Rep_induct = thm Rep_induct, Abs_induct = thm Abs_induct}) end; structure Data = Generic_Data ( type T = info list Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge_list (K true) data; ); fun get_info_generic context = Symtab.lookup_list (Data.get context) #> map (transform_info (Morphism.transfer_morphism'' context)); val get_info = get_info_generic o Context.Proof; val get_info_global = get_info_generic o Context.Theory; fun put_info name info = Data.map (Symtab.cons_list (name, transform_info Morphism.trim_context_morphism info)); (* global interpretation *) structure Typedef_Plugin = Plugin(type T = string); val typedef_plugin = Plugin_Name.declare_setup \<^binding>\typedef\; fun interpretation f = Typedef_Plugin.interpretation typedef_plugin (fn name => fn lthy => lthy |> Local_Theory.map_background_naming (Name_Space.root_path #> Name_Space.add_path (Long_Name.qualifier name)) |> f name |> Local_Theory.restore_background_naming lthy); (* primitive typedef axiomatization -- for fresh typedecl *) val typedef_overloaded = Attrib.setup_config_bool \<^binding>\typedef_overloaded\ (K false); fun mk_inhabited A = let val T = HOLogic.dest_setT (Term.fastype_of A) in HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, HOLogic.mk_mem (Bound 0, A))) end; fun mk_typedef newT oldT RepC AbsC A = let val typedefC = Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end; fun primitive_typedef {overloaded} type_definition_name newT oldT Rep_name Abs_name A lthy = let (* errors *) fun show_names pairs = commas_quote (map fst pairs); val lhs_tfrees = Term.add_tfreesT newT []; val rhs_tfrees = Term.add_tfreesT oldT []; val _ = (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => () | extras => error ("Extra type variables in representing set: " ^ show_names extras)); val _ = (case Term.add_frees A [] of [] => [] | xs => error ("Illegal variables in representing set: " ^ show_names xs)); (* axiomatization *) val ((RepC, AbsC), consts_lthy) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((Rep_name, newT --> oldT), NoSyn) ##>> Sign.declare_const lthy ((Abs_name, oldT --> newT), NoSyn)); val const_dep = Theory.const_dep (Proof_Context.theory_of consts_lthy); val defs_context = Proof_Context.defs_context consts_lthy; val A_consts = fold_aterms (fn Const c => insert (op =) (const_dep c) | _ => I) A []; val A_types = (fold_types o fold_subtypes) (fn Type t => insert (op =) (Theory.type_dep t) | _ => I) A []; val typedef_deps = A_consts @ A_types; val newT_dep = Theory.type_dep (dest_Type newT); val ((axiom_name, axiom), axiom_lthy) = consts_lthy |> Local_Theory.background_theory_result (Thm.add_axiom consts_lthy (type_definition_name, mk_typedef newT oldT RepC AbsC A) ##> Theory.add_deps defs_context "" newT_dep typedef_deps ##> Theory.add_deps defs_context "" (const_dep (dest_Const RepC)) [newT_dep] ##> Theory.add_deps defs_context "" (const_dep (dest_Const AbsC)) [newT_dep]); val axiom_defs = Theory.defs_of (Proof_Context.theory_of axiom_lthy); val newT_deps = maps #2 (Defs.get_deps axiom_defs (#1 newT_dep)); val _ = if null newT_deps orelse overloaded orelse Config.get lthy typedef_overloaded then () else error (Pretty.string_of (Pretty.chunks [Pretty.paragraph (Pretty.text "Type definition with open dependencies, use" @ [Pretty.brk 1, Pretty.str "\"", Pretty.keyword1 "typedef", Pretty.brk 1, Pretty.str "(", Pretty.keyword2 "overloaded", Pretty.str ")\"", Pretty.brk 1] @ Pretty.text "or enable configuration option \"typedef_overloaded\" in the context."), Pretty.block [Pretty.str " Type:", Pretty.brk 2, Syntax.pretty_typ axiom_lthy newT], Pretty.block (Pretty.str " Deps:" :: Pretty.brk 2 :: Pretty.commas (map (Defs.pretty_entry (Proof_Context.defs_context axiom_lthy)) newT_deps))])) in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end; (* derived bindings *) type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding}; fun prefix_binding prfx name = Binding.reset_pos (Binding.qualify_name false name (prfx ^ Binding.name_of name)); fun qualify_binding name = Binding.qualify false (Binding.name_of name); fun default_bindings name = {Rep_name = prefix_binding "Rep_" name, Abs_name = prefix_binding "Abs_" name, type_definition_name = prefix_binding "type_definition_" name}; fun make_bindings name NONE = default_bindings name | make_bindings _ (SOME bindings) = bindings; fun make_morphisms name NONE = default_bindings name | make_morphisms name (SOME (Rep_name, Abs_name)) = {Rep_name = qualify_binding name Rep_name, Abs_name = qualify_binding name Abs_name, type_definition_name = #type_definition_name (default_bindings name)}; (* prepare_typedef *) fun prepare_typedef prep_term overloaded (name, raw_args, mx) raw_set opt_bindings lthy = let (* rhs *) val tmp_ctxt = lthy |> fold (Variable.declare_typ o TFree) raw_args; val set = prep_term tmp_ctxt raw_set; val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set; val setT = Term.fastype_of set; val oldT = HOLogic.dest_setT setT handle TYPE _ => error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT)); val bname = Binding.name_of name; val goal = mk_inhabited set; val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Lexicon.read_variable bname), setT)); (* lhs *) val args = map (Proof_Context.check_tfree tmp_ctxt') raw_args; val (newT, typedecl_lthy) = lthy |> Typedecl.typedecl {final = false} (name, args, mx) ||> Variable.declare_term set; val Type (full_name, _) = newT; (* axiomatization *) val {Rep_name, Abs_name, type_definition_name} = make_bindings name opt_bindings; val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = typedecl_lthy |> primitive_typedef overloaded type_definition_name newT oldT Rep_name Abs_name set; val alias_lthy = typedef_lthy |> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC)) |> Local_Theory.const_alias Abs_name (#1 (Term.dest_Const AbsC)); (* result *) fun note ((b, atts), th) = Local_Theory.note ((b, atts), [th]) #>> (fn (_, [th']) => th'); fun typedef_result inhabited lthy1 = let val ((_, [type_definition]), lthy2) = lthy1 |> Local_Theory.note ((type_definition_name, []), [inhabited RS typedef]); fun make th = Goal.norm_result lthy2 (type_definition RS th); val (((((((((Rep, Rep_inverse), Abs_inverse), Rep_inject), Abs_inject), Rep_cases), Abs_cases), Rep_induct), Abs_induct), lthy3) = lthy2 |> note ((Rep_name, []), make @{thm type_definition.Rep}) ||>> note ((Binding.suffix_name "_inverse" Rep_name, []), make @{thm type_definition.Rep_inverse}) ||>> note ((Binding.suffix_name "_inverse" Abs_name, []), make @{thm type_definition.Abs_inverse}) ||>> note ((Binding.suffix_name "_inject" Rep_name, []), make @{thm type_definition.Rep_inject}) ||>> note ((Binding.suffix_name "_inject" Abs_name, []), make @{thm type_definition.Abs_inject}) ||>> note ((Binding.suffix_name "_cases" Rep_name, [Attrib.case_names [Binding.name_of Rep_name], - Attrib.internal (K (Induct.cases_pred full_name))]), + Attrib.internal \<^here> (K (Induct.cases_pred full_name))]), make @{thm type_definition.Rep_cases}) ||>> note ((Binding.suffix_name "_cases" Abs_name, [Attrib.case_names [Binding.name_of Abs_name], - Attrib.internal (K (Induct.cases_type full_name))]), + Attrib.internal \<^here> (K (Induct.cases_type full_name))]), make @{thm type_definition.Abs_cases}) ||>> note ((Binding.suffix_name "_induct" Rep_name, [Attrib.case_names [Binding.name_of Rep_name], - Attrib.internal (K (Induct.induct_pred full_name))]), + Attrib.internal \<^here> (K (Induct.induct_pred full_name))]), make @{thm type_definition.Rep_induct}) ||>> note ((Binding.suffix_name "_induct" Abs_name, [Attrib.case_names [Binding.name_of Abs_name], - Attrib.internal (K (Induct.induct_type full_name))]), + Attrib.internal \<^here> (K (Induct.induct_type full_name))]), make @{thm type_definition.Abs_induct}); val info = ({rep_type = oldT, abs_type = newT, Rep_name = #1 (Term.dest_Const RepC), Abs_name = #1 (Term.dest_Const AbsC), axiom_name = axiom_name}, {inhabited = inhabited, type_definition = type_definition, Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse, Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases, Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct}); in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => put_info full_name (transform_info phi info)) |> Typedef_Plugin.data Plugin_Name.default_filter full_name |> pair (full_name, info) end; in ((goal, goal_pat, typedef_result), alias_lthy) end handle ERROR msg => cat_error msg ("The error(s) above occurred in typedef " ^ Binding.print name); (* add_typedef: tactic interface *) fun add_typedef overloaded typ set opt_bindings tac lthy = let val ((goal, _, typedef_result), lthy') = prepare_typedef Syntax.check_term overloaded typ set opt_bindings lthy; val inhabited = Goal.prove lthy' [] [] goal (tac o #context) |> Goal.norm_result lthy'; in typedef_result inhabited lthy' end; fun add_typedef_global overloaded typ set opt_bindings tac = Named_Target.theory_map_result (apsnd o transform_info) (add_typedef overloaded typ set opt_bindings tac) (* typedef: proof interface *) local fun gen_typedef prep_term prep_constraint overloaded (b, raw_args, mx) set opt_bindings lthy = let val args = map (apsnd (prep_constraint lthy)) raw_args; val ((goal, goal_pat, typedef_result), lthy') = prepare_typedef prep_term overloaded (b, args, mx) set opt_bindings lthy; fun after_qed [[th]] = snd o typedef_result th; in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end; in val typedef = gen_typedef Syntax.check_term (K I); val typedef_cmd = gen_typedef Syntax.read_term Typedecl.read_constraint; end; (** outer syntax **) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\typedef\ "HOL type definition (requires non-emptiness proof)" (Parse_Spec.overloaded -- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- Parse.term) -- Scan.option (\<^keyword>\morphisms\ |-- Parse.!!! (Parse.binding -- Parse.binding)) >> (fn (((((overloaded, vs), t), mx), A), opt_morphs) => fn lthy => typedef_cmd {overloaded = overloaded} (t, vs, mx) A (SOME (make_morphisms t opt_morphs)) lthy)); val overloaded = typedef_overloaded; (** theory export **) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parent_spaces = map Sign.type_space (Theory.parents_of thy); val typedefs = Name_Space.dest_table (#types (Type.rep_tsig (Sign.tsig_of thy))) |> maps (fn (name, _) => if exists (fn space => Name_Space.declared space name) parent_spaces then [] else get_info_global thy name |> map (fn ({rep_type, abs_type, Rep_name, Abs_name, axiom_name}, _) => (name, (rep_type, (abs_type, (Rep_name, (Abs_name, axiom_name))))))); val encode = let open XML.Encode Term_XML.Encode in list (pair string (pair typ (pair typ (pair string (pair string string))))) end; in if null typedefs then () else Export_Theory.export_body thy "typedefs" (encode typedefs) end else ()); end; diff --git a/src/Provers/order_tac.ML b/src/Provers/order_tac.ML --- a/src/Provers/order_tac.ML +++ b/src/Provers/order_tac.ML @@ -1,434 +1,435 @@ signature REIFY_TABLE = sig type table val empty : table val get_var : term -> table -> (int * table) val get_term : int -> table -> term option end structure Reifytab: REIFY_TABLE = struct type table = (int * (term * int) list) val empty = (0, []) fun get_var t (max_var, tis) = (case AList.lookup Envir.aeconv tis t of SOME v => (v, (max_var, tis)) | NONE => (max_var, (max_var + 1, (t, max_var) :: tis)) ) fun get_term v (_, tis) = Library.find_first (fn (_, v2) => v = v2) tis |> Option.map fst end signature LOGIC_SIGNATURE = sig val mk_Trueprop : term -> term val dest_Trueprop : term -> term val Trueprop_conv : conv -> conv val Not : term val conj : term val disj : term val notI : thm (* (P \ False) \ \ P *) val ccontr : thm (* (\ P \ False) \ P *) val conjI : thm (* P \ Q \ P \ Q *) val conjE : thm (* P \ Q \ (P \ Q \ R) \ R *) val disjE : thm (* P \ Q \ (P \ R) \ (Q \ R) \ R *) val not_not_conv : conv (* \ (\ P) \ P *) val de_Morgan_conj_conv : conv (* \ (P \ Q) \ \ P \ \ Q *) val de_Morgan_disj_conv : conv (* \ (P \ Q) \ \ P \ \ Q *) val conj_disj_distribL_conv : conv (* P \ (Q \ R) \ (P \ Q) \ (P \ R) *) val conj_disj_distribR_conv : conv (* (Q \ R) \ P \ (Q \ P) \ (R \ P) *) end (* Control tracing output of the solver. *) val order_trace_cfg = Attrib.setup_config_bool @{binding "order_trace"} (K false) (* In partial orders, literals of the form \ x < y will force the order solver to perform case distinctions, which leads to an exponential blowup of the runtime. The split limit controls the number of literals of this form that are passed to the solver. *) val order_split_limit_cfg = Attrib.setup_config_int @{binding "order_split_limit"} (K 8) datatype order_kind = Order | Linorder type order_literal = (bool * Order_Procedure.order_atom) type order_context = { kind : order_kind, ops : term list, thms : (string * thm) list, conv_thms : (string * thm) list } signature BASE_ORDER_TAC = sig val tac : (order_literal Order_Procedure.fm -> Order_Procedure.prf_trm option) -> order_context -> thm list -> Proof.context -> int -> tactic end functor Base_Order_Tac( structure Logic_Sig : LOGIC_SIGNATURE; val excluded_types : typ list) : BASE_ORDER_TAC = struct open Order_Procedure fun expect _ (SOME x) = x | expect f NONE = f () fun list_curry0 f = (fn [] => f, 0) fun list_curry1 f = (fn [x] => f x, 1) fun list_curry2 f = (fn [x, y] => f x y, 2) fun dereify_term consts reifytab t = let fun dereify_term' (App (t1, t2)) = (dereify_term' t1) $ (dereify_term' t2) | dereify_term' (Const s) = AList.lookup (op =) consts s |> expect (fn () => raise TERM ("Const " ^ s ^ " not in", map snd consts)) | dereify_term' (Var v) = Reifytab.get_term (integer_of_int v) reifytab |> the in dereify_term' t end fun dereify_order_fm (eq, le, lt) reifytab t = let val consts = [ ("eq", eq), ("le", le), ("lt", lt), ("Not", Logic_Sig.Not), ("disj", Logic_Sig.disj), ("conj", Logic_Sig.conj) ] in dereify_term consts reifytab t end fun strip_AppP t = let fun strip (AppP (f, s), ss) = strip (f, s::ss) | strip x = x in strip (t, []) end fun replay_conv convs cvp = let val convs = convs @ [("all_conv", list_curry0 Conv.all_conv)] @ map (apsnd list_curry1) [ ("atom_conv", I), ("neg_atom_conv", I), ("arg_conv", Conv.arg_conv)] @ map (apsnd list_curry2) [ ("combination_conv", Conv.combination_conv), ("then_conv", curry (op then_conv))] fun lookup_conv convs c = AList.lookup (op =) convs c |> expect (fn () => error ("Can't replay conversion: " ^ c)) fun rp_conv t = (case strip_AppP t ||> map rp_conv of (PThm c, cvs) => let val (conv, arity) = lookup_conv convs c in if arity = length cvs then conv cvs else error ("Expected " ^ Int.toString arity ^ " arguments for conversion " ^ c ^ " but got " ^ (length cvs |> Int.toString) ^ " arguments") end | _ => error "Unexpected constructor in conversion proof") in rp_conv cvp end fun replay_prf_trm replay_conv dereify ctxt thmtab assmtab p = let fun replay_prf_trm' _ (PThm s) = AList.lookup (op =) thmtab s |> expect (fn () => error ("Cannot replay theorem: " ^ s)) | replay_prf_trm' assmtab (Appt (p, t)) = replay_prf_trm' assmtab p |> Drule.infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (dereify t))] | replay_prf_trm' assmtab (AppP (p1, p2)) = apply2 (replay_prf_trm' assmtab) (p2, p1) |> op COMP | replay_prf_trm' assmtab (AbsP (reified_t, p)) = let val t = dereify reified_t val t_thm = Logic_Sig.mk_Trueprop t |> Thm.cterm_of ctxt |> Assumption.assume ctxt val rp = replay_prf_trm' (Termtab.update (Thm.prop_of t_thm, t_thm) assmtab) p in Thm.implies_intr (Thm.cprop_of t_thm) rp end | replay_prf_trm' assmtab (Bound reified_t) = let val t = dereify reified_t |> Logic_Sig.mk_Trueprop in Termtab.lookup assmtab t |> expect (fn () => raise TERM ("Assumption not found:", t::Termtab.keys assmtab)) end | replay_prf_trm' assmtab (Conv (t, cp, p)) = let val thm = replay_prf_trm' assmtab (Bound t) val conv = Logic_Sig.Trueprop_conv (replay_conv cp) val conv_thm = Conv.fconv_rule conv thm val conv_term = Thm.prop_of conv_thm in replay_prf_trm' (Termtab.update (conv_term, conv_thm) assmtab) p end in replay_prf_trm' assmtab p end fun replay_order_prf_trm ord_ops {thms = thms, conv_thms = conv_thms, ...} ctxt reifytab assmtab = let val thmtab = thms @ [ ("conjE", Logic_Sig.conjE), ("conjI", Logic_Sig.conjI), ("disjE", Logic_Sig.disjE) ] val convs = map (apsnd list_curry0) ( map (apsnd Conv.rewr_conv) conv_thms @ [ ("not_not_conv", Logic_Sig.not_not_conv), ("de_Morgan_conj_conv", Logic_Sig.de_Morgan_conj_conv), ("de_Morgan_disj_conv", Logic_Sig.de_Morgan_disj_conv), ("conj_disj_distribR_conv", Logic_Sig.conj_disj_distribR_conv), ("conj_disj_distribL_conv", Logic_Sig.conj_disj_distribL_conv) ]) val dereify = dereify_order_fm ord_ops reifytab in replay_prf_trm (replay_conv convs) dereify ctxt thmtab assmtab end fun strip_Not (nt $ t) = if nt = Logic_Sig.Not then t else nt $ t | strip_Not t = t fun limit_not_less [_, _, lt] ctxt decomp_prems = let val thy = Proof_Context.theory_of ctxt val trace = Config.get ctxt order_trace_cfg val limit = Config.get ctxt order_split_limit_cfg fun is_not_less_term t = case try (strip_Not o Logic_Sig.dest_Trueprop) t of SOME (binop $ _ $ _) => Pattern.matches thy (lt, binop) | NONE => false val not_less_prems = filter (is_not_less_term o Thm.prop_of o fst) decomp_prems val _ = if trace andalso length not_less_prems > limit then tracing "order split limit exceeded" else () in filter_out (is_not_less_term o Thm.prop_of o fst) decomp_prems @ take limit not_less_prems end fun decomp [eq, le, lt] ctxt t = let fun is_excluded t = exists (fn ty => ty = fastype_of t) excluded_types fun decomp'' (binop $ t1 $ t2) = let open Order_Procedure val thy = Proof_Context.theory_of ctxt fun try_match pat = try (Pattern.match thy (pat, binop)) (Vartab.empty, Vartab.empty) in if is_excluded t1 then NONE else case (try_match eq, try_match le, try_match lt) of (SOME env, _, _) => SOME (true, EQ, (t1, t2), env) | (_, SOME env, _) => SOME (true, LEQ, (t1, t2), env) | (_, _, SOME env) => SOME (true, LESS, (t1, t2), env) | _ => NONE end | decomp'' _ = NONE fun decomp' (nt $ t) = if nt = Logic_Sig.Not then decomp'' t |> Option.map (fn (b, c, p, e) => (not b, c, p, e)) else decomp'' (nt $ t) | decomp' t = decomp'' t in try Logic_Sig.dest_Trueprop t |> Option.mapPartial decomp' end fun maximal_envs envs = let fun test_opt p (SOME x) = p x | test_opt _ NONE = false fun leq_env (tyenv1, tenv1) (tyenv2, tenv2) = Vartab.forall (fn (v, ty) => Vartab.lookup tyenv2 v |> test_opt (fn ty2 => ty2 = ty)) tyenv1 andalso Vartab.forall (fn (v, (ty, t)) => Vartab.lookup tenv2 v |> test_opt (fn (ty2, t2) => ty2 = ty andalso t2 aconv t)) tenv1 fun fold_env (i, env) es = fold_index (fn (i2, env2) => fn es => if i = i2 then es else if leq_env env env2 then (i, i2) :: es else es) envs es val env_order = fold_index fold_env envs [] val graph = fold_index (fn (i, env) => fn g => Int_Graph.new_node (i, env) g) envs Int_Graph.empty val graph = fold Int_Graph.add_edge env_order graph val strong_conns = Int_Graph.strong_conn graph val maximals = filter (fn comp => length comp = length (Int_Graph.all_succs graph comp)) strong_conns in map (Int_Graph.all_preds graph) maximals end fun order_tac raw_order_proc octxt simp_prems = Subgoal.FOCUS (fn {prems=prems, context=ctxt, ...} => let val trace = Config.get ctxt order_trace_cfg fun these' _ [] = [] | these' f (x :: xs) = case f x of NONE => these' f xs | SOME y => (x, y) :: these' f xs val prems = simp_prems @ prems |> filter (fn p => null (Term.add_vars (Thm.prop_of p) [])) |> map (Conv.fconv_rule Thm.eta_conversion) val decomp_prems = these' (decomp (#ops octxt) ctxt o Thm.prop_of) prems fun env_of (_, (_, _, _, env)) = env val env_groups = maximal_envs (map env_of decomp_prems) fun order_tac' (_, []) = no_tac | order_tac' (env, decomp_prems) = let val [eq, le, lt] = #ops octxt |> map (Envir.eta_contract o Envir.subst_term env) val decomp_prems = case #kind octxt of Order => limit_not_less (#ops octxt) ctxt decomp_prems | _ => decomp_prems fun reify_prem (_, (b, ctor, (x, y), _)) (ps, reifytab) = (Reifytab.get_var x ##>> Reifytab.get_var y) reifytab |>> (fn vp => (b, ctor (apply2 Int_of_integer vp)) :: ps) val (reified_prems, reifytab) = fold_rev reify_prem decomp_prems ([], Reifytab.empty) val reified_prems_conj = foldl1 (fn (x, a) => And (x, a)) (map Atom reified_prems) val prems_conj_thm = map fst decomp_prems |> foldl1 (fn (x, a) => Logic_Sig.conjI OF [x, a]) |> Conv.fconv_rule Thm.eta_conversion val prems_conj = prems_conj_thm |> Thm.prop_of val proof = raw_order_proc reified_prems_conj val pretty_term_list = Pretty.list "" "" o map (Syntax.pretty_term (Config.put show_types true ctxt)) val pretty_thm_list = Pretty.list "" "" o map (Thm.pretty_thm ctxt) fun pretty_type_of t = Pretty.block [ Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt (Term.fastype_of t)) ] fun pretty_trace () = [ ("order kind:", Pretty.str (@{make_string} (#kind octxt))) , ("order operators:", Pretty.block [ pretty_term_list [eq, le, lt], Pretty.brk 1 , pretty_type_of le ]) , ("premises:", pretty_thm_list prems) , ("selected premises:", pretty_thm_list (map fst decomp_prems)) , ("reified premises:", Pretty.str (@{make_string} reified_prems)) , ("contradiction:", Pretty.str (@{make_string} (Option.isSome proof))) ] |> map (fn (t, pp) => Pretty.block [Pretty.str t, Pretty.brk 1, pp]) |> Pretty.big_list "order solver called with the parameters" val _ = if trace then tracing (Pretty.string_of (pretty_trace ())) else () val assmtab = Termtab.make [(prems_conj, prems_conj_thm)] val replay = replay_order_prf_trm (eq, le, lt) octxt ctxt reifytab assmtab in case proof of NONE => no_tac | SOME p => SOLVED' (resolve_tac ctxt [replay p]) 1 end in map (fn is => ` (env_of o hd) (map (nth decomp_prems) is) |> order_tac') env_groups |> FIRST end) val ad_absurdum_tac = SUBGOAL (fn (A, i) => case try (Logic_Sig.dest_Trueprop o Logic.strip_assums_concl) A of SOME (nt $ _) => if nt = Logic_Sig.Not then resolve0_tac [Logic_Sig.notI] i else resolve0_tac [Logic_Sig.ccontr] i | _ => resolve0_tac [Logic_Sig.ccontr] i) fun tac raw_order_proc octxt simp_prems ctxt = ad_absurdum_tac THEN' order_tac raw_order_proc octxt simp_prems ctxt end functor Order_Tac(structure Base_Tac : BASE_ORDER_TAC) = struct fun order_context_eq ({kind = kind1, ops = ops1, ...}, {kind = kind2, ops = ops2, ...}) = kind1 = kind2 andalso eq_list (op aconv) (ops1, ops2) fun order_data_eq (x, y) = order_context_eq (fst x, fst y) structure Data = Generic_Data( type T = (order_context * (order_context -> thm list -> Proof.context -> int -> tactic)) list val empty = [] fun merge data = Library.merge order_data_eq data ) fun declare (octxt as {kind = kind, raw_proc = raw_proc, ...}) lthy = - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val ops = map (Morphism.term phi) (#ops octxt) - val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt) - val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt) - val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms} - in - context |> Data.map (Library.insert order_data_eq (octxt', raw_proc)) - end) + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} + (fn phi => fn context => + let + val ops = map (Morphism.term phi) (#ops octxt) + val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt) + val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt) + val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms} + in + context |> Data.map (Library.insert order_data_eq (octxt', raw_proc)) + end) fun declare_order { ops = {eq = eq, le = le, lt = lt}, thms = { trans = trans, (* x \ y \ y \ z \ x \ z *) refl = refl, (* x \ x *) eqD1 = eqD1, (* x = y \ x \ y *) eqD2 = eqD2, (* x = y \ y \ x *) antisym = antisym, (* x \ y \ y \ x \ x = y *) contr = contr (* \ P \ P \ R *) }, conv_thms = { less_le = less_le, (* x < y \ x \ y \ x \ y *) nless_le = nless_le (* \ a < b \ \ a \ b \ a = b *) } } = declare { kind = Order, ops = [eq, le, lt], thms = [("trans", trans), ("refl", refl), ("eqD1", eqD1), ("eqD2", eqD2), ("antisym", antisym), ("contr", contr)], conv_thms = [("less_le", less_le), ("nless_le", nless_le)], raw_proc = Base_Tac.tac Order_Procedure.po_contr_prf } fun declare_linorder { ops = {eq = eq, le = le, lt = lt}, thms = { trans = trans, (* x \ y \ y \ z \ x \ z *) refl = refl, (* x \ x *) eqD1 = eqD1, (* x = y \ x \ y *) eqD2 = eqD2, (* x = y \ y \ x *) antisym = antisym, (* x \ y \ y \ x \ x = y *) contr = contr (* \ P \ P \ R *) }, conv_thms = { less_le = less_le, (* x < y \ x \ y \ x \ y *) nless_le = nless_le, (* \ x < y \ y \ x *) nle_le = nle_le (* \ a \ b \ b \ a \ b \ a *) } } = declare { kind = Linorder, ops = [eq, le, lt], thms = [("trans", trans), ("refl", refl), ("eqD1", eqD1), ("eqD2", eqD2), ("antisym", antisym), ("contr", contr)], conv_thms = [("less_le", less_le), ("nless_le", nless_le), ("nle_le", nle_le)], raw_proc = Base_Tac.tac Order_Procedure.lo_contr_prf } (* Try to solve the goal by calling the order solver with each of the declared orders. *) fun tac simp_prems ctxt = let fun app_tac (octxt, tac0) = CHANGED o tac0 octxt simp_prems ctxt in FIRST' (map app_tac (Data.get (Context.Proof ctxt))) end end diff --git a/src/Pure/Isar/attrib.ML b/src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML +++ b/src/Pure/Isar/attrib.ML @@ -1,658 +1,660 @@ (* Title: Pure/Isar/attrib.ML Author: Markus Wenzel, TU Muenchen Symbolic representation of attributes -- with name and syntax. *) signature ATTRIB = sig type thms = Attrib.thms type fact = Attrib.fact val print_attributes: bool -> Proof.context -> unit val attribute_space: Context.generic -> Name_Space.T val define_global: binding -> (Token.src -> attribute) -> string -> theory -> string * theory val define: binding -> (Token.src -> attribute) -> string -> local_theory -> string * local_theory val check_name_generic: Context.generic -> xstring * Position.T -> string val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val attribs: Token.src list context_parser val opt_attribs: Token.src list context_parser val pretty_attribs: Proof.context -> Token.src list -> Pretty.T list val pretty_binding: Proof.context -> Attrib.binding -> string -> Pretty.T list val attribute: Proof.context -> Token.src -> attribute val attribute_global: theory -> Token.src -> attribute val attribute_cmd: Proof.context -> Token.src -> attribute val attribute_cmd_global: theory -> Token.src -> attribute val map_specs: ('a list -> 'att list) -> (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list val map_facts: ('a list -> 'att list) -> (('c * 'a list) * ('d * 'a list) list) list -> (('c * 'att list) * ('d * 'att list) list) list val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) -> (('c * 'a list) * ('b * 'a list) list) list -> (('c * 'att list) * ('fact * 'att list) list) list val trim_context_binding: Attrib.binding -> Attrib.binding val trim_context_thms: thms -> thms val trim_context_fact: fact -> fact val global_notes: string -> fact list -> theory -> (string * thm list) list * theory val local_notes: string -> fact list -> Proof.context -> (string * thm list) list * Proof.context val generic_notes: string -> fact list -> Context.generic -> (string * thm list) list * Context.generic val lazy_notes: string -> binding * thm list lazy -> Context.generic -> Context.generic val eval_thms: Proof.context -> (Facts.ref * Token.src list) list -> thm list val attribute_syntax: attribute context_parser -> Token.src -> attribute val setup: binding -> attribute context_parser -> string -> theory -> theory val local_setup: binding -> attribute context_parser -> string -> local_theory -> local_theory val attribute_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory - val internal: (morphism -> attribute) -> Token.src - val internal_declaration: Morphism.declaration_entity -> thms + val internal: Position.T -> (morphism -> attribute) -> Token.src + val internal_declaration: Position.T -> Morphism.declaration_entity -> thms val add_del: attribute -> attribute -> attribute context_parser val thm: thm context_parser val thms: thm list context_parser val multi_thm: thm list context_parser val transform_facts: morphism -> fact list -> fact list val partial_evaluation: Proof.context -> fact list -> fact list val print_options: bool -> Proof.context -> unit val config_bool: binding -> (Context.generic -> bool) -> bool Config.T * (theory -> theory) val config_int: binding -> (Context.generic -> int) -> int Config.T * (theory -> theory) val config_real: binding -> (Context.generic -> real) -> real Config.T * (theory -> theory) val config_string: binding -> (Context.generic -> string) -> string Config.T * (theory -> theory) val setup_config_bool: binding -> (Context.generic -> bool) -> bool Config.T val setup_config_int: binding -> (Context.generic -> int) -> int Config.T val setup_config_real: binding -> (Context.generic -> real) -> real Config.T val setup_config_string: binding -> (Context.generic -> string) -> string Config.T val option_bool: string * Position.T -> bool Config.T * (theory -> theory) val option_int: string * Position.T -> int Config.T * (theory -> theory) val option_real: string * Position.T -> real Config.T * (theory -> theory) val option_string: string * Position.T -> string Config.T * (theory -> theory) val setup_option_bool: string * Position.T -> bool Config.T val setup_option_int: string * Position.T -> int Config.T val setup_option_real: string * Position.T -> real Config.T val setup_option_string: string * Position.T -> string Config.T val consumes: int -> Token.src val constraints: int -> Token.src val cases_open: Token.src val case_names: string list -> Token.src val case_conclusion: string * string list -> Token.src end; structure Attrib: sig type binding = Attrib.binding include ATTRIB end = struct open Attrib; (** named attributes **) (* theory data *) structure Attributes = Generic_Data ( type T = ((Token.src -> attribute) * string) Name_Space.table; val empty : T = Name_Space.empty_table Markup.attributeN; fun merge data : T = Name_Space.merge_tables data; ); val ops_attributes = {get_data = Attributes.get, put_data = Attributes.put}; val get_attributes = Attributes.get o Context.Proof; fun print_attributes verbose ctxt = let val attribs = get_attributes ctxt; fun prt_attr (name, (_, "")) = Pretty.mark_str name | prt_attr (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "attributes:" (map prt_attr (Name_Space.markup_table verbose ctxt attribs))] |> Pretty.writeln_chunks end; val attribute_space = Name_Space.space_of_table o Attributes.get; (* define *) fun define_global binding att comment = Entity.define_global ops_attributes binding (att, comment); fun define binding att comment = Entity.define ops_attributes binding (att, comment); (* check *) fun check_name_generic context = #1 o Name_Space.check context (Attributes.get context); val check_name = check_name_generic o Context.Proof; fun check_src ctxt src = let val _ = if Token.checked_src src then () else Context_Position.report ctxt (#1 (Token.range_of src)) Markup.language_attribute; in #1 (Token.check_src ctxt get_attributes src) end; val attribs = Args.context -- Scan.lift Parse.attribs >> (fn (ctxt, srcs) => map (check_src ctxt) srcs); val opt_attribs = Scan.optional attribs []; (* pretty printing *) fun pretty_attribs _ [] = [] | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)]; fun pretty_binding ctxt (b, atts) sep = (case (Binding.is_empty b, null atts) of (true, true) => [] | (false, true) => [Pretty.block [Binding.pretty b, Pretty.str sep]] | (true, false) => [Pretty.block (pretty_attribs ctxt atts @ [Pretty.str sep])] | (false, false) => [Pretty.block (Binding.pretty b :: Pretty.brk 1 :: pretty_attribs ctxt atts @ [Pretty.str sep])]); (* get attributes *) fun attribute_generic context = let val table = Attributes.get context in fn src => let val name = #1 (Token.name_of_src src); val label = Long_Name.qualify Markup.attributeN name; val att = #1 (Name_Space.get table name) src; in Position.setmp_thread_data_label label att : attribute end end; val attribute = attribute_generic o Context.Proof; val attribute_global = attribute_generic o Context.Theory; fun attribute_cmd ctxt = attribute ctxt o check_src ctxt; fun attribute_cmd_global thy = attribute_global thy o check_src (Proof_Context.init_global thy); (* attributed declarations *) fun map_specs f = map (apfst (apsnd f)); fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g))); (* implicit context *) val trim_context_binding: Attrib.binding -> Attrib.binding = apsnd ((map o map) Token.trim_context); val trim_context_thms: thms -> thms = map (fn (thms, atts) => (map Thm.trim_context thms, (map o map) Token.trim_context atts)); fun trim_context_fact (binding, thms) = (trim_context_binding binding, trim_context_thms thms); (* fact expressions *) fun global_notes kind facts thy = thy |> Global_Theory.note_thmss kind (map_facts (map (attribute_global thy)) facts); fun local_notes kind facts ctxt = ctxt |> Proof_Context.note_thmss kind (map_facts (map (attribute ctxt)) facts); fun generic_notes kind facts context = context |> Context.mapping_result (global_notes kind facts) (local_notes kind facts); fun lazy_notes kind arg = Context.mapping (Global_Theory.add_thms_lazy kind arg) (Proof_Context.add_thms_lazy kind arg); fun eval_thms ctxt srcs = ctxt |> Proof_Context.note_thmss "" (map_facts_refs (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) [(Binding.empty_atts, srcs)]) |> fst |> maps snd; (* attribute setup *) fun attribute_syntax scan src (context, th) = let val (a, context') = Token.syntax_generic scan src context in a (context', th) end; fun setup binding scan comment = define_global binding (attribute_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (attribute_syntax scan) comment #> snd; fun attribute_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Attrib.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (") " ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* internal attribute *) fun make_name ctxt name = Token.make_src (name, Position.none) [] |> check_src ctxt |> hd; local val internal_binding = Binding.make ("attribute", \<^here>); val _ = Theory.setup (setup internal_binding (Scan.lift Args.internal_attribute >> Morphism.form || Scan.lift Args.internal_declaration >> (Thm.declaration_attribute o K o Morphism.form)) "internal attribute"); val internal_name = make_name (Context.the_local_context ()) (Binding.name_of internal_binding); -val internal_arg = Token.make_string0 ""; -fun internal_source value = [internal_name, Token.assign (SOME value) internal_arg]; +fun internal_source name value = [internal_name, Token.assign (SOME value) (Token.make_string name)]; in -fun internal arg = internal_source (Token.Attribute (Morphism.entity arg)); -fun internal_declaration arg = [([Drule.dummy_thm], [internal_source (Token.Declaration arg)])]; +fun internal pos arg = + internal_source ("", pos) (Token.Attribute (Morphism.entity arg)); + +fun internal_declaration pos arg = + [([Drule.dummy_thm], [internal_source ("", pos) (Token.Declaration arg)])]; end; (* add/del syntax *) fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); (** parsing attributed theorems **) local val fact_name = Parse.position Args.internal_fact >> (fn (_, pos) => ("", pos)) || Args.name_position; fun gen_thm pick = Scan.depend (fn context => let val get = Proof_Context.get_fact_generic context; val get_fact = get o Facts.Fact; fun get_named is_sel pos name = let val (a, ths) = get (Facts.Named ((name, pos), NONE)) in (if is_sel then NONE else a, ths) end; in Parse.$$$ "[" |-- Scan.pass context attribs --| Parse.$$$ "]" >> (fn srcs => let val atts = map (attribute_generic context) srcs; val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); in (context', pick ("", Position.none) [th']) end) || (Scan.ahead Args.alt_name -- Args.named_fact get_fact >> (fn (s, fact) => ("", Facts.Fact s, fact)) || Scan.ahead (fact_name -- Scan.option Parse.thm_sel) :|-- (fn ((name, pos), sel) => Args.named_fact (get_named (is_some sel) pos) --| Scan.option Parse.thm_sel >> (fn fact => (name, Facts.Named ((name, pos), sel), fact)))) -- Scan.pass context opt_attribs >> (fn ((name, thmref, fact), srcs) => let val ths = Facts.select thmref fact; val atts = map (attribute_generic context) srcs; val (ths', context') = fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; in (context', pick (name, Facts.ref_pos thmref) ths') end) end); in val thm = gen_thm Facts.the_single; val multi_thm = gen_thm (K I); val thms = Scan.repeats multi_thm; end; (* transform fact expressions *) fun transform_facts phi = map (fn ((a, atts), bs) => ((Morphism.binding phi a, (map o map) (Token.transform phi) atts), bs |> map (fn (ths, btts) => (Morphism.fact phi ths, (map o map) (Token.transform phi) btts)))); (** partial evaluation -- observing rule/declaration/mixed attributes **) (*NB: result length may change due to rearrangement of symbolic expression*) local fun apply_att src (context, th) = let val src1 = map Token.init_assignable src; val result = attribute_generic context src1 (context, th); val src2 = map Token.closure src1; in (src2, result) end; fun err msg src = let val (name, pos) = Token.name_of_src src in error (msg ^ " " ^ quote name ^ Position.here pos) end; fun eval src ((th, dyn), (decls, context)) = (case (apply_att src (context, th), dyn) of ((_, (NONE, SOME th')), NONE) => ((th', NONE), (decls, context)) | ((_, (NONE, SOME _)), SOME _) => err "Mixed dynamic attribute followed by static rule" src | ((src', (SOME context', NONE)), NONE) => let val decls' = (case decls of [] => [(th, [src'])] | (th2, srcs2) :: rest => if Thm.eq_thm_strict (th, th2) then ((th2, src' :: srcs2) :: rest) else (th, [src']) :: (th2, srcs2) :: rest); in ((th, NONE), (decls', context')) end | ((src', (opt_context', opt_th')), _) => let val context' = the_default context opt_context'; val th' = the_default th opt_th'; val dyn' = (case dyn of NONE => SOME (th, [src']) | SOME (dyn_th, srcs) => SOME (dyn_th, src' :: srcs)); in ((th', dyn'), (decls, context')) end); in fun partial_evaluation ctxt facts = (facts, Context.Proof (Context_Position.not_really ctxt)) |-> fold_map (fn ((b, more_atts), fact) => fn context => let val (fact', (decls, context')) = (fact, ([], context)) |-> fold_map (fn (ths, atts) => fn res1 => (ths, res1) |-> fold_map (fn th => fn res2 => let val ((th', dyn'), res3) = fold eval (atts @ more_atts) ((th, NONE), res2); val th_atts' = (case dyn' of NONE => (th', []) | SOME (dyn_th', atts') => (dyn_th', rev atts')); in (th_atts', res3) end)) |>> flat; val decls' = rev (map (apsnd rev) decls); val facts' = if eq_list (eq_fst Thm.eq_thm_strict) (decls', fact') then [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] else if null decls' then [((b, []), fact')] else [(Binding.empty_atts, decls'), ((b, []), fact')]; in (facts', context') end) |> fst |> flat |> map (apsnd (map (apfst single))) |> filter_out (fn (b, fact) => Binding.is_empty_atts b andalso forall (null o #2) fact); end; (** configuration options **) (* naming *) structure Configs = Theory_Data ( type T = Config.value Config.T Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ); fun print_options verbose ctxt = let fun prt (name, config) = let val value = Config.get ctxt config in Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="), Pretty.brk 1, Pretty.str (Config.print_value value)] end; val space = attribute_space (Context.Proof ctxt); val configs = Name_Space.markup_entries verbose ctxt space (Symtab.dest (Configs.get (Proof_Context.theory_of ctxt))); in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; (* concrete syntax *) local val equals = Parse.$$$ "="; fun scan_value (Config.Bool _) = equals -- Args.$$$ "false" >> K (Config.Bool false) || equals -- Args.$$$ "true" >> K (Config.Bool true) || Scan.succeed (Config.Bool true) | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | scan_value (Config.String _) = equals |-- Args.name >> Config.String; fun scan_config thy config = let val config_type = Config.get_global thy config in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end; fun register binding config thy = let val name = Sign.full_name thy binding in thy |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form_entity) "configuration option" |> Configs.map (Symtab.update (name, config)) end; fun declare make coerce binding default = let val name = Binding.name_of binding; val pos = Binding.pos_of binding; val config_value = Config.declare (name, pos) (make o default); val config = coerce config_value; in (config, register binding config_value) end; in fun register_config config = register (Binding.make (Config.name_of config, Config.pos_of config)) config; val register_config_bool = register_config o Config.bool_value; val register_config_int = register_config o Config.int_value; val register_config_real = register_config o Config.real_value; val register_config_string = register_config o Config.string_value; val config_bool = declare Config.Bool Config.bool; val config_int = declare Config.Int Config.int; val config_real = declare Config.Real Config.real; val config_string = declare Config.String Config.string; end; (* implicit setup *) local fun setup_config declare_config binding default = let val (config, setup) = declare_config binding default; val _ = Theory.setup setup; in config end; in val setup_config_bool = setup_config config_bool; val setup_config_int = setup_config config_int; val setup_config_string = setup_config config_string; val setup_config_real = setup_config config_real; end; (* system options *) local fun declare_option coerce (name, pos) = let val config = Config.declare_option (name, pos); in (coerce config, register_config config) end; fun setup_option coerce (name, pos) = let val config = Config.declare_option (name, pos); val _ = Theory.setup (register_config config); in coerce config end; in val option_bool = declare_option Config.bool; val option_int = declare_option Config.int; val option_real = declare_option Config.real; val option_string = declare_option Config.string; val setup_option_bool = setup_option Config.bool; val setup_option_int = setup_option Config.int; val setup_option_real = setup_option Config.real; val setup_option_string = setup_option Config.string; end; (* theory setup *) val _ = Theory.setup (setup \<^binding>\tagged\ (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #> setup \<^binding>\untagged\ (Scan.lift Args.name >> Thm.untag) "untagged theorem" #> setup \<^binding>\kind\ (Scan.lift Args.name >> Thm.kind) "theorem kind" #> setup \<^binding>\THEN\ (Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))) "resolution with rule" #> setup \<^binding>\OF\ (thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))) "rule resolved with facts" #> setup \<^binding>\rename_abs\ (Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs => Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))) "rename bound variables in abstractions" #> setup \<^binding>\unfolded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))) "unfolded definitions" #> setup \<^binding>\folded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))) "folded definitions" #> setup \<^binding>\consumes\ (Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes) "number of consumed facts" #> setup \<^binding>\constraints\ (Scan.lift Parse.nat >> Rule_Cases.constraints) "number of equality constraints" #> setup \<^binding>\cases_open\ (Scan.succeed Rule_Cases.cases_open) "rule with open parameters" #> setup \<^binding>\case_names\ (Scan.lift (Scan.repeat (Args.name -- Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") [])) >> (fn cs => Rule_Cases.cases_hyp_names (map #1 cs) (map (map (the_default Rule_Cases.case_hypsN) o #2) cs))) "named rule cases" #> setup \<^binding>\case_conclusion\ (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) "named conclusion of rule cases" #> setup \<^binding>\params\ (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) "named rule parameters" #> setup \<^binding>\rule_format\ (Scan.lift (Args.mode "no_asm") >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)) "result put into canonical rule format" #> setup \<^binding>\elim_format\ (Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))) "destruct rule turned into elimination rule format" #> setup \<^binding>\no_vars\ (Scan.succeed (Thm.rule_attribute [] (Variable.import_vars o Context.proof_of))) "imported schematic variables" #> setup \<^binding>\atomize\ (Scan.succeed Object_Logic.declare_atomize) "declaration of atomize rule" #> setup \<^binding>\rulify\ (Scan.succeed Object_Logic.declare_rulify) "declaration of rulify rule" #> setup \<^binding>\rotated\ (Scan.lift (Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))) "rotated theorem premises" #> setup \<^binding>\defn\ (add_del Local_Defs.defn_add Local_Defs.defn_del) "declaration of definitional transformations" #> setup \<^binding>\abs_def\ (Scan.succeed (Thm.rule_attribute [] (Local_Defs.abs_def_rule o Context.proof_of))) "abstract over free variables of definitional theorem" #> register_config_bool Goal.quick_and_dirty #> register_config_bool Ast.trace #> register_config_bool Ast.stats #> register_config_bool Printer.show_brackets #> register_config_bool Printer.show_sorts #> register_config_bool Printer.show_types #> register_config_bool Printer.show_markup #> register_config_bool Printer.show_structs #> register_config_bool Printer.show_question_marks #> register_config_bool Syntax.ambiguity_warning #> register_config_int Syntax.ambiguity_limit #> register_config_bool Syntax_Trans.eta_contract #> register_config_bool Name_Space.names_long #> register_config_bool Name_Space.names_short #> register_config_bool Name_Space.names_unique #> register_config_int ML_Print_Depth.print_depth #> register_config_string ML_Env.ML_environment #> register_config_bool ML_Env.ML_read_global #> register_config_bool ML_Env.ML_write_global #> register_config_bool ML_Options.source_trace #> register_config_bool ML_Options.exception_trace #> register_config_bool ML_Options.exception_debugger #> register_config_bool ML_Options.debugger #> register_config_bool Proof_Context.show_abbrevs #> register_config_int Goal_Display.goals_limit #> register_config_bool Goal_Display.show_main_goal #> register_config_bool Thm.show_consts #> register_config_bool Thm.show_hyps #> register_config_bool Thm.show_tags #> register_config_bool Pattern.unify_trace_failure #> register_config_int Unify.trace_bound #> register_config_int Unify.search_bound #> register_config_bool Unify.trace_simp #> register_config_bool Unify.trace_types #> register_config_int Raw_Simplifier.simp_depth_limit #> register_config_int Raw_Simplifier.simp_trace_depth_limit #> register_config_bool Raw_Simplifier.simp_debug #> register_config_bool Raw_Simplifier.simp_trace #> register_config_bool Local_Defs.unfold_abs_def); (* internal source *) local val make_name0 = make_name (Context.the_local_context ()); val consumes_name = make_name0 "consumes"; val constraints_name = make_name0 "constraints"; val cases_open_name = make_name0 "cases_open"; val case_names_name = make_name0 "case_names"; val case_conclusion_name = make_name0 "case_conclusion"; in fun consumes i = consumes_name :: Token.make_int i; fun constraints i = constraints_name :: Token.make_int i; val cases_open = [cases_open_name]; fun case_names names = case_names_name :: map Token.make_string0 names; fun case_conclusion (name, names) = case_conclusion_name :: map Token.make_string0 (name :: names); end; end; \ No newline at end of file diff --git a/src/Pure/Isar/bundle.ML b/src/Pure/Isar/bundle.ML --- a/src/Pure/Isar/bundle.ML +++ b/src/Pure/Isar/bundle.ML @@ -1,251 +1,251 @@ (* Title: Pure/Isar/bundle.ML Author: Makarius Bundled declarations (notes etc.). *) signature BUNDLE = sig type name = string val bundle_space: Context.generic -> Name_Space.T val check: Proof.context -> xstring * Position.T -> string val get: Proof.context -> name -> Attrib.thms val read: Proof.context -> xstring * Position.T -> Attrib.thms val extern: Proof.context -> string -> xstring val print_bundles: bool -> Proof.context -> unit val bundle: binding * Attrib.thms -> (binding * typ option * mixfix) list -> local_theory -> local_theory val bundle_cmd: binding * (Facts.ref * Token.src list) list -> (binding * string option * mixfix) list -> local_theory -> local_theory val init: binding -> theory -> local_theory val unbundle: name list -> local_theory -> local_theory val unbundle_cmd: (xstring * Position.T) list -> local_theory -> local_theory val includes: name list -> Proof.context -> Proof.context val includes_cmd: (xstring * Position.T) list -> Proof.context -> Proof.context val include_: name list -> Proof.state -> Proof.state val include_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state val including: name list -> Proof.state -> Proof.state val including_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state end; structure Bundle: BUNDLE = struct (** context data **) structure Data = Generic_Data ( type T = Attrib.thms Name_Space.table * Attrib.thms option; val empty : T = (Name_Space.empty_table Markup.bundleN, NONE); fun merge ((tab1, target1), (tab2, target2)) = (Name_Space.merge_tables (tab1, tab2), merge_options (target1, target2)); ); (* bundles *) type name = string val get_all_generic = #1 o Data.get; val get_all = get_all_generic o Context.Proof; val bundle_space = Name_Space.space_of_table o #1 o Data.get; fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_all ctxt); val get = Name_Space.get o get_all_generic o Context.Proof; fun read ctxt = get ctxt o check ctxt; fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt)); (* target -- bundle under construction *) fun the_target thy = #2 (Data.get (Context.Theory thy)) |> \<^if_none>\error "Missing bundle target"\; val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE; val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms; fun augment_target thms = Local_Theory.background_theory (fn thy => set_target (the_target thy @ thms) thy); (* print bundles *) fun pretty_bundle ctxt (markup_name, bundle) = let val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt; fun prt_thm_attribs atts th = Pretty.block (Pretty.breaks (prt_thm th :: Attrib.pretty_attribs ctxt atts)); fun prt_thms (ths, []) = map prt_thm ths | prt_thms (ths, atts) = map (prt_thm_attribs atts) ths; in Pretty.block ([Pretty.keyword1 "bundle", Pretty.str " ", Pretty.mark_str markup_name] @ (if null bundle then [] else Pretty.fbreaks (Pretty.str " =" :: maps prt_thms bundle))) end; fun print_bundles verbose ctxt = Pretty.writeln_chunks (map (pretty_bundle ctxt) (Name_Space.markup_table verbose ctxt (get_all ctxt))); (** define bundle **) (* context and morphisms *) val trim_context_bundle = map (fn (fact, atts) => (map Thm.trim_context fact, (map o map) Token.trim_context atts)); fun transfer_bundle thy = map (fn (fact, atts) => (map (Thm.transfer thy) fact, (map o map) (Token.transfer thy) atts)); fun transform_bundle phi = map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts)); fun define_bundle (b, bundle) context = let val (name, bundles') = get_all_generic context |> Name_Space.define context true (b, trim_context_bundle bundle); val context' = (Data.map o apfst o K) bundles' context; in (name, context') end; (* command *) local fun gen_bundle prep_fact prep_att add_fixes (binding, raw_bundle) raw_fixes lthy = let val (_, ctxt') = add_fixes raw_fixes lthy; val bundle0 = raw_bundle |> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts)); val bundle = Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)] |> maps #2 |> transform_bundle (Proof_Context.export_morphism ctxt' lthy) |> trim_context_bundle; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi in #2 (define_bundle (Morphism.binding psi binding, transform_bundle psi bundle) context) end) end; in val bundle = gen_bundle (K I) (K I) Proof_Context.add_fixes; val bundle_cmd = gen_bundle Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; end; (* target *) local fun bad_operation _ = error "Not possible in bundle target"; fun conclude invisible binding = Local_Theory.background_theory_result (fn thy => thy |> invisible ? Context_Position.set_visible_global false |> Context.Theory |> define_bundle (binding, the_target thy) ||> (Context.the_theory #> invisible ? Context_Position.restore_visible_global thy #> reset_target)); fun pretty binding lthy = let val bundle = the_target (Proof_Context.theory_of lthy); val (name, lthy') = lthy |> Local_Theory.raw_theory (Context_Position.set_visible_global false) |> conclude true binding; val thy_ctxt' = Proof_Context.init_global (Proof_Context.theory_of lthy'); val markup_name = Name_Space.markup_extern thy_ctxt' (Name_Space.space_of_table (get_all thy_ctxt')) name; in [pretty_bundle lthy' (markup_name, bundle)] end; fun bundle_notes kind facts lthy = let val bundle = facts |> maps (fn ((_, more_atts), thms) => map (fn (ths, atts) => (ths, atts @ more_atts)) thms); in lthy |> augment_target (transform_bundle (Local_Theory.standard_morphism_theory lthy) bundle) |> Generic_Target.standard_notes (op <>) kind facts |> Attrib.local_notes kind facts end; -fun bundle_declaration decl lthy = +fun bundle_declaration pos decl lthy = lthy - |> (augment_target o Attrib.internal_declaration) + |> (augment_target o Attrib.internal_declaration pos) (Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl) |> Generic_Target.standard_declaration (K true) decl; in fun init binding thy = thy |> Local_Theory.init {background_naming = Sign.naming_of thy, setup = set_target [] #> Proof_Context.init_global, conclude = conclude false binding #> #2} {define = bad_operation, notes = bundle_notes, abbrev = bad_operation, - declaration = K bundle_declaration, + declaration = fn {pos, ...} => bundle_declaration pos, theory_registration = bad_operation, locale_dependency = bad_operation, - pretty = pretty binding} + pretty = pretty binding}; end; (** activate bundles **) local fun gen_activate notes prep_bundle args ctxt = let val thy = Proof_Context.theory_of ctxt; val decls = maps (prep_bundle ctxt) args |> transfer_bundle thy; in ctxt |> Context_Position.set_visible false |> notes [(Binding.empty_atts, decls)] |> #2 |> Context_Position.restore_visible ctxt end; fun gen_unbundle prep_bundle = gen_activate Local_Theory.notes prep_bundle; fun gen_includes prep_bundle = gen_activate (Attrib.local_notes "") prep_bundle; fun gen_include prep_bundle bs = Proof.assert_forward #> Proof.map_context (gen_includes prep_bundle bs) #> Proof.reset_facts; fun gen_including prep_bundle bs = Proof.assert_backward #> Proof.map_context (gen_includes prep_bundle bs) in val unbundle = gen_unbundle get; val unbundle_cmd = gen_unbundle read; val includes = gen_includes get; val includes_cmd = gen_includes read; val include_ = gen_include get; val include_cmd = gen_include read; val including = gen_including get; val including_cmd = gen_including read; end; end; diff --git a/src/Pure/Isar/code.ML b/src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML +++ b/src/Pure/Isar/code.ML @@ -1,1570 +1,1570 @@ (* Title: Pure/Isar/code.ML Author: Florian Haftmann, TU Muenchen Abstract executable ingredients of theory. Management of data dependent on executable ingredients as synchronized cache; purged on any change of underlying executable ingredients. *) signature CODE = sig (*constants*) val check_const: theory -> term -> string val read_const: theory -> string -> string val string_of_const: theory -> string -> string val args_number: theory -> string -> int (*constructor sets*) val constrset_of_consts: theory -> (string * typ) list -> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list) (*code equations and certificates*) val assert_eqn: theory -> thm * bool -> thm * bool val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string) type cert val constrain_cert: theory -> sort list -> cert -> cert val conclude_cert: cert -> cert val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list val equations_of_cert: theory -> cert -> ((string * sort) list * typ) * (((string option * term) list * (string option * term)) * (thm option * bool)) list option val pretty_cert: theory -> cert -> Pretty.T list (*executable code*) type constructors type abs_type val type_interpretation: (string -> theory -> theory) -> theory -> theory val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory val declare_datatype_global: (string * typ) list -> theory -> theory val declare_datatype_cmd: string list -> theory -> theory val declare_abstype: thm -> local_theory -> local_theory val declare_abstype_global: thm -> theory -> theory val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory val declare_default_eqns_global: (thm * bool) list -> theory -> theory val declare_eqns: (thm * bool) list -> local_theory -> local_theory val declare_eqns_global: (thm * bool) list -> theory -> theory val add_eqn_global: thm * bool -> theory -> theory val del_eqn_global: thm -> theory -> theory val declare_abstract_eqn: thm -> local_theory -> local_theory val declare_abstract_eqn_global: thm -> theory -> theory val declare_aborting_global: string -> theory -> theory val declare_unimplemented_global: string -> theory -> theory val declare_case_global: thm -> theory -> theory val declare_undefined_global: string -> theory -> theory val get_type: theory -> string -> constructors * bool val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option val is_constr: theory -> string -> bool val is_abstr: theory -> string -> bool val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list -> string -> cert type case_schema val get_case_schema: theory -> string -> case_schema option val get_case_cong: theory -> string -> thm option val is_undefined: theory -> string -> bool val print_codesetup: theory -> unit end; signature CODE_DATA_ARGS = sig type T val empty: T end; signature CODE_DATA = sig type T val change: theory option -> (T -> T) -> T val change_yield: theory option -> (T -> 'a * T) -> 'a * T end; signature PRIVATE_CODE = sig include CODE val declare_data: Any.T -> serial val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a) -> theory -> ('a -> 'b * 'a) -> 'b * 'a end; structure Code : PRIVATE_CODE = struct (** auxiliary **) (* printing *) fun string_of_typ thy = Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); fun string_of_const thy c = let val ctxt = Proof_Context.init_global thy in case Axclass.inst_of_param thy c of SOME (c, tyco) => Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]" (Proof_Context.extern_type ctxt tyco) | NONE => Proof_Context.extern_const ctxt c end; (* constants *) fun const_typ thy = Type.strip_sorts o Sign.the_const_type thy; fun args_number thy = length o binder_types o const_typ thy; fun devarify ty = let val tys = build (fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty); val vs = Name.invent Name.context Name.aT (length tys); val mapping = map2 (fn v => fn (vi, sort) => (vi, TFree (v, sort))) vs tys; in Term.typ_subst_TVars mapping ty end; fun typscheme thy (c, ty) = (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); fun typscheme_equiv (ty1, ty2) = Type.raw_instance (devarify ty1, ty2) andalso Type.raw_instance (devarify ty2, ty1); fun check_bare_const thy t = case try dest_Const t of SOME c_ty => c_ty | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); fun check_unoverload thy (c, ty) = let val c' = Axclass.unoverload_const thy (c, ty); val ty_decl = const_typ thy c'; in if typscheme_equiv (ty_decl, Logic.varifyT_global ty) then c' else error ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_const thy = check_unoverload thy o check_bare_const thy; fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; fun read_const thy = check_unoverload thy o read_bare_const thy; (** executable specifications **) (* types *) datatype type_spec = Constructors of { constructors: (string * ((string * sort) list * typ list)) list, case_combinators: string list} | Abstractor of { abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string, more_abstract_functions: string list}; fun concrete_constructors_of (Constructors {constructors, ...}) = constructors | concrete_constructors_of _ = []; fun constructors_of (Constructors {constructors, ...}) = (constructors, false) | constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) = ([(co, (vs, [ty]))], true); fun case_combinators_of (Constructors {case_combinators, ...}) = case_combinators | case_combinators_of (Abstractor _) = []; fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) = (vs, Constructors {constructors = constructors, case_combinators = insert (op =) c case_combinators}); fun projection_of (Constructors _) = NONE | projection_of (Abstractor {projection, ...}) = SOME projection; fun abstract_functions_of (Constructors _) = [] | abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) = projection :: more_abstract_functions; fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) = (vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = insert (op =) c more_abstract_functions}); fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1}, Constructors {case_combinators = case_combinators2, ...}) = Constructors {constructors = constructors, case_combinators = merge (op =) (case_combinators1, case_combinators2)} | join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1}, Abstractor {more_abstract_functions = more_abstract_functions2, ...}) = Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)}; fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2)); (* functions *) datatype fun_spec = Eqns of bool * (thm * bool) list | Proj of term * (string * string) | Abstr of thm * (string * string); val unimplemented = Eqns (true, []); fun is_unimplemented (Eqns (true, [])) = true | is_unimplemented _ = false; fun is_default (Eqns (true, _)) = true | is_default _ = false; val aborting = Eqns (false, []); fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs | associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs | associated_abstype _ = NONE; (* cases *) type case_schema = int * (int * (string * int) option list); datatype case_spec = No_Case | Case of {schema: case_schema, tycos: string list, cong: thm} | Undefined; fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map fst (map_filter I raw_cos)) | associated_datatypes _ = ([], []); (** background theory data store **) (* historized declaration data *) structure History = struct type 'a T = { entry: 'a, suppressed: bool, (*incompatible entries are merely suppressed after theory merge but sustain*) history: serial list (*explicit trace of declaration history supports non-monotonic declarations*) } Symtab.table; fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry | some_entry _ = NONE; fun lookup table = Symtab.lookup table #> some_entry; fun register key entry table = if is_some (Symtab.lookup table key) then Symtab.map_entry key (fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table; fun modify_entry key f = Symtab.map_entry key (fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history}); fun all table = Symtab.dest table |> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE); local fun merge_history join_same ({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) = let val history = merge (op =) (history1, history2); val entry = if hd history1 = hd history2 then join_same (entry1, entry2) else if hd history = hd history1 then entry1 else entry2; in {entry = entry, suppressed = false, history = history} end; in fun join join_same tables = Symtab.join (K (merge_history join_same)) tables; fun suppress key = Symtab.map_entry key (fn {entry, history, ...} => {entry = entry, suppressed = true, history = history}); fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} => {entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history}); end; end; datatype specs = Specs of { types: ((string * sort) list * type_spec) History.T, pending_eqns: (thm * bool) list Symtab.table, functions: fun_spec History.T, cases: case_spec History.T }; fun types_of (Specs {types, ...}) = types; fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns; fun functions_of (Specs {functions, ...}) = functions; fun cases_of (Specs {cases, ...}) = cases; fun make_specs (types, ((pending_eqns, functions), cases)) = Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}; val empty_specs = make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty)); fun map_specs f (Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}) = make_specs (f (types, ((pending_eqns, functions), cases))); fun merge_specs (Specs {types = types1, pending_eqns = _, functions = functions1, cases = cases1}, Specs {types = types2, pending_eqns = _, functions = functions2, cases = cases2}) = let val types = History.join join_same_types (types1, types2); val all_types = map (snd o snd) (History.all types); fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of NONE => true | SOME (tyco, abs) => (case History.lookup types tyco of NONE => false | SOME (_, Constructors _) => false | SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) => abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c)); fun check_datatypes (_, case_spec) = let val (tycos, required_constructors) = associated_datatypes case_spec; val allowed_constructors = tycos |> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types) |> map fst; in subset (op =) (required_constructors, allowed_constructors) end; val all_constructors = maps (fst o constructors_of) all_types; val functions = History.join fst (functions1, functions2) |> fold (History.suppress o fst) all_constructors |> History.suppress_except check_abstype; val cases = History.join fst (cases1, cases2) |> History.suppress_except check_datatypes; in make_specs (types, ((Symtab.empty, functions), cases)) end; val map_types = map_specs o apfst; val map_pending_eqns = map_specs o apsnd o apfst o apfst; val map_functions = map_specs o apsnd o apfst o apsnd; val map_cases = map_specs o apsnd o apsnd; (* data slots dependent on executable code *) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); local type kind = {empty: Any.T}; val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table); fun invoke f k = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => f kind | NONE => raise Fail "Invalid code data identifier"); in fun declare_data empty = let val k = serial (); val kind = {empty = empty}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; fun invoke_init k = invoke (fn kind => #empty kind) k; end; (*local*) (* global theory store *) local type data = Any.T Datatab.table; fun make_dataref () = Synchronized.var "code data" (NONE : (data * Context.theory_id) option); structure Code_Data = Theory_Data ( type T = specs * (string * (data * Context.theory_id) option Synchronized.var); val empty = (empty_specs, (Context.theory_long_name (Context.the_global_context ()), make_dataref ())); fun merge ((specs1, dataref), (specs2, _)) = (merge_specs (specs1, specs2), dataref); ); fun init_dataref thy = let val thy_name = Context.theory_long_name thy in if #1 (#2 (Code_Data.get thy)) = thy_name then NONE else SOME ((Code_Data.map o apsnd) (K (thy_name, make_dataref ())) thy) end; in val _ = Theory.setup (Theory.at_begin init_dataref); (* access to executable specifications *) val specs_of : theory -> specs = fst o Code_Data.get; fun modify_specs f thy = let val thy_name = Context.theory_long_name thy in Code_Data.map (fn (specs, _) => (f specs, (thy_name, make_dataref ()))) thy end; (* access to data dependent on executable specifications *) fun change_yield_data (kind, mk, dest) theory f = let val dataref = #2 (#2 (Code_Data.get theory)); val (datatab, thy_id) = case Synchronized.value dataref of SOME (datatab, thy_id) => if Context.eq_thy_id (Context.theory_id theory, thy_id) then (datatab, thy_id) else (Datatab.empty, Context.theory_id theory) | NONE => (Datatab.empty, Context.theory_id theory) val data = case Datatab.lookup datatab kind of SOME data => data | NONE => invoke_init kind; val result as (_, data') = f (dest data); val _ = Synchronized.change dataref ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id)); in result end; end; (*local*) (* pending function equations *) (* Ideally, *all* equations implementing a functions would be treated as *one* atomic declaration; unfortunately, we cannot implement this: the too-well-established declaration interface are Isar attributes which operate on *one* single theorem. Hence we treat such Isar declarations as "pending" and historize them as proper declarations at the end of each theory. *) fun modify_pending_eqns c f specs = let val existing_eqns = case History.lookup (functions_of specs) c of SOME (Eqns (false, eqns)) => eqns | _ => []; in specs |> map_pending_eqns (Symtab.map_default (c, existing_eqns) f) end; fun register_fun_spec c spec = map_pending_eqns (Symtab.delete_safe c) #> map_functions (History.register c spec); fun lookup_fun_spec specs c = case Symtab.lookup (pending_eqns_of specs) c of SOME eqns => Eqns (false, eqns) | NONE => (case History.lookup (functions_of specs) c of SOME spec => spec | NONE => unimplemented); fun lookup_proper_fun_spec specs c = let val spec = lookup_fun_spec specs c in if is_unimplemented spec then NONE else SOME spec end; fun all_fun_specs specs = map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c)) (union (op =) ((Symtab.keys o pending_eqns_of) specs) ((Symtab.keys o functions_of) specs)); fun historize_pending_fun_specs thy = let val pending_eqns = (pending_eqns_of o specs_of) thy; in if Symtab.is_empty pending_eqns then NONE else thy |> modify_specs (map_functions (Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns) #> map_pending_eqns (K Symtab.empty)) |> SOME end; val _ = Theory.setup (Theory.at_end historize_pending_fun_specs); (** foundation **) (* types *) fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s); fun analyze_constructor thy (c, ty) = let val _ = Thm.global_cterm_of thy (Const (c, ty)); val ty_decl = devarify (const_typ thy c); fun last_typ c_ty ty = let val tfrees = Term.add_tfreesT ty []; val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty)) handle TYPE _ => no_constr thy "bad type" c_ty val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else (); val _ = if has_duplicates (eq_fst (op =)) vs then no_constr thy "duplicate type variables in datatype" c_ty else (); val _ = if length tfrees <> length vs then no_constr thy "type variables missing in datatype" c_ty else (); in (tyco, vs) end; val (tyco, _) = last_typ (c, ty) ty_decl; val (_, vs) = last_typ (c, ty) ty; in ((tyco, map snd vs), (c, (map fst vs, ty))) end; fun constrset_of_consts thy consts = let val _ = map (fn (c, _) => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) consts; val raw_constructors = map (analyze_constructor thy) consts; val tyco = case distinct (op =) (map (fst o fst) raw_constructors) of [tyco] => tyco | [] => error "Empty constructor set" | tycos => error ("Different type constructors in constructor set: " ^ commas_quote tycos) val vs = Name.invent Name.context Name.aT (Sign.arity_number thy tyco); fun inst vs' (c, (vs, ty)) = let val the_v = the o AList.lookup (op =) (vs ~~ vs'); val ty' = map_type_tfree (fn (v, _) => TFree (the_v v, [])) ty; val (vs'', ty'') = typscheme thy (c, ty'); in (c, (vs'', binder_types ty'')) end; val constructors = map (inst vs o snd) raw_constructors; in (tyco, (map (rpair []) vs, constructors)) end; fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy); type constructors = (string * sort) list * (string * ((string * sort) list * typ list)) list; fun get_type thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec) | NONE => Sign.arity_number thy tyco |> Name.invent Name.context Name.aT |> map (rpair []) |> rpair [] |> rpair false; type abs_type = (string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string}; fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) => (vs, {abs_rep = Thm.transfer thy abs_rep, abstractor = abstractor, projection = projection}) | _ => error ("Not an abstract type: " ^ tyco); fun get_type_of_constr_or_abstr thy c = case (body_type o const_typ thy) c of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end | _ => NONE; fun is_constr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, false) => true | _ => false; fun is_abstr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, true) => true | _ => false; (* bare code equations *) (* convention for variables: ?x ?'a for free-floating theorems (e.g. in the data store) ?x 'a for certificates x 'a for final representation of equations *) exception BAD_THM of string; fun bad_thm msg = raise BAD_THM msg; datatype strictness = Silent | Liberal | Strict fun handle_strictness thm_of f strictness thy x = SOME (f x) handle BAD_THM msg => case strictness of Silent => NONE | Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE) | Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); fun is_linear thm = let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm in not (has_duplicates (op =) ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])) end; fun check_decl_ty thy (c, ty) = let val ty_decl = const_typ thy c; in if typscheme_equiv (ty_decl, ty) then () else bad_thm ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) = let fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | Free _ => bad_thm "Illegal free variable" | _ => I) t []; fun tvars_of t = fold_term_types (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v | TFree _ => bad_thm "Illegal free type variable")) t []; val lhs_vs = vars_of lhs; val rhs_vs = vars_of rhs; val lhs_tvs = tvars_of lhs; val rhs_tvs = tvars_of rhs; val _ = if null (subtract (op =) lhs_vs rhs_vs) then () else bad_thm "Free variables on right hand side of equation"; val _ = if null (subtract (op =) lhs_tvs rhs_tvs) then () else bad_thm "Free type variables on right hand side of equation"; val (head, args) = strip_comb lhs; val (c, ty) = case head of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty) | _ => bad_thm "Equation not headed by constant"; fun check _ (Abs _) = bad_thm "Abstraction on left hand side of equation" | check 0 (Var _) = () | check _ (Var _) = bad_thm "Variable with application on left hand side of equation" | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | check n (Const (c_ty as (c, ty))) = if allow_pats then let val c' = Axclass.unoverload_const thy c_ty in if n = (length o binder_types) ty then if allow_consts orelse is_constr thy c' then () else bad_thm (quote c ^ " is not a constructor, on left hand side of equation") else bad_thm ("Partially applied constant " ^ quote c ^ " on left hand side of equation") end else bad_thm ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation") val _ = map (check 0) args; val _ = if allow_nonlinear orelse is_linear thm then () else bad_thm "Duplicate variables on left hand side of equation"; val _ = if (is_none o Axclass.class_of_param thy) c then () else bad_thm "Overloaded constant as head in equation"; val _ = if not (is_constr thy c) then () else bad_thm "Constructor as head in equation"; val _ = if not (is_abstr thy c) then () else bad_thm "Abstractor as head in equation"; val _ = check_decl_ty thy (c, ty); val _ = case strip_type ty of (Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of SOME (_, type_spec) => (case projection_of type_spec of SOME proj => if c = proj then bad_thm "Projection as head in equation" else () | _ => ()) | _ => ()) | _ => (); in () end; local fun raw_assert_eqn thy check_patterns (thm, proper) = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val _ = check_eqn thy {allow_nonlinear = not proper, allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs); in (thm, proper) end; fun raw_assert_abs_eqn thy some_tyco thm = let val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val (proj_t, lhs) = dest_comb full_lhs handle TERM _ => bad_thm "Not an abstract equation"; val (proj, ty) = dest_Const proj_t handle TERM _ => bad_thm "Not an abstract equation"; val (tyco, Ts) = (dest_Type o domain_type) ty handle TERM _ => bad_thm "Not an abstract equation" | TYPE _ => bad_thm "Not an abstract equation"; val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () else bad_thm ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco') | NONE => (); val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec) | _ => bad_thm ("Not an abstract type: " ^ tyco); val _ = if proj = proj' then () else bad_thm ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj'); val _ = check_eqn thy {allow_nonlinear = false, allow_consts = false, allow_pats = false} thm (lhs, rhs); val _ = if ListPair.all (fn (T, (_, sort)) => Sign.of_sort thy (T, sort)) (Ts, vs) then () else error ("Type arguments do not satisfy sort constraints of abstype certificate."); in (thm, (tyco, abs')) end; in fun generic_assert_eqn strictness thy check_patterns eqn = handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn; fun generic_assert_abs_eqn strictness thy check_patterns thm = handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm; end; fun assert_eqn thy = the o generic_assert_eqn Strict thy true; fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco; val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun const_typ_eqn thy thm = let val (c, ty) = head_eqn thm; val c' = Axclass.unoverload_const thy (c, ty); (*permissive wrt. to overloaded constants!*) in (c', ty) end; fun const_eqn thy = fst o const_typ_eqn thy; fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun mk_proj tyco vs ty abs rep = let val ty_abs = Type (tyco, map TFree vs); val xarg = Var (("x", 0), ty); in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; (* technical transformations of code equations *) fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); fun same_arity thy thms = let val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms; val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0; fun expand_eta (lhs, rhs) thm = let val l = k - length (snd (strip_comb lhs)); val (raw_vars, _) = Term.strip_abs_eta l rhs; val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) raw_vars; fun expand (v, ty) thm = Drule.fun_cong_rule thm (Thm.global_cterm_of thy (Var ((v, 0), ty))); in thm |> fold expand vars |> Conv.fconv_rule Drule.beta_eta_conversion end; in map2 expand_eta lhs_rhss thms end; fun mk_desymbolization pre post mk vs = let val names = map (pre o fst o fst) vs |> map (Name.desymbolize (SOME false)) |> Name.variant_list [] |> map post; in map_filter (fn (((v, i), x), v') => if v = v' andalso i = 0 then NONE else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) end; fun desymbolize_tvars thy thms = let val tvs = build (fold (Term.add_tvars o Thm.prop_of) thms); val instT = mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs; in map (Thm.instantiate (TVars.make instT, Vars.empty)) thms end; fun desymbolize_vars thy thm = let val vs = Term.add_vars (Thm.prop_of thm) []; val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs; in Thm.instantiate (TVars.empty, Vars.make inst) thm end; fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); (* preparation and classification of code equations *) fun prep_eqn strictness thy = apfst (meta_rewrite thy) #> generic_assert_eqn strictness thy false #> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn)); fun prep_eqns strictness thy = map_filter (prep_eqn strictness thy) #> AList.group (op =); fun prep_abs_eqn strictness thy = meta_rewrite thy #> generic_assert_abs_eqn strictness thy NONE #> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn)); fun prep_maybe_abs_eqn thy raw_thm = let val thm = meta_rewrite thy raw_thm; val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm; in case some_abs_thm of SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco)) | NONE => generic_assert_eqn Liberal thy false (thm, false) |> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE))) end; (* abstype certificates *) local fun raw_abstype_cert thy proto_thm = let val thm = (Axclass.unoverload (Proof_Context.init_global thy) o meta_rewrite thy) proto_thm; val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb) o apfst dest_Const o dest_comb) lhs handle TERM _ => bad_thm "Not an abstype certificate"; val _ = apply2 (fn c => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep); val _ = check_decl_ty thy (abs, raw_ty); val _ = check_decl_ty thy (rep, rep_ty); val _ = if length (binder_types raw_ty) = 1 then () else bad_thm "Bad type for abstract constructor"; val _ = (fst o dest_Var) param handle TERM _ => bad_thm "Not an abstype certificate"; val _ = if param = rhs then () else bad_thm "Not an abstype certificate"; val ((tyco, sorts), (abs, (vs, ty'))) = analyze_constructor thy (abs, devarify raw_ty); val ty = domain_type ty'; val (vs', _) = typscheme thy (abs, ty'); in (tyco, (vs ~~ sorts, ((abs, (vs', ty)), (rep, thm)))) end; in fun check_abstype_cert strictness thy proto_thm = handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm; end; (* code equation certificates *) fun build_head thy (c, ty) = Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); fun get_head thy cert_thm = let val [head] = Thm.chyps_of cert_thm; val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; in (typscheme thy (c, ty), head) end; fun typscheme_projection thy = typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; fun typscheme_abs thy = typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of; fun constrain_thm thy vs sorts thm = let val mapping = map2 (fn (v, sort) => fn sort' => (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; val instT = TVars.build (fold2 (fn (v, sort) => fn (_, sort') => TVars.add (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping); val subst = (Term.map_types o map_type_tfree) (fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); in thm |> Thm.varifyT_global |> Thm.instantiate (instT, Vars.empty) |> pair subst end; fun concretify_abs thy tyco abs_thm = let val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco; val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm val ty = fastype_of lhs; val ty_abs = (fastype_of o snd o dest_comb) lhs; val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs)); val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm]; in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; fun add_rhss_of_eqn thy t = let val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t; fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) | add_const _ = I val add_consts = fold_aterms add_const in add_consts rhs o fold add_consts args end; val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global; abstype cert = Nothing of thm | Equations of thm * bool list | Projection of term * string | Abstract of thm * string with fun dummy_thm ctxt c = let val thy = Proof_Context.theory_of ctxt; val raw_ty = devarify (const_typ thy c); val (vs, _) = typscheme thy (c, raw_ty); val sortargs = case Axclass.class_of_param thy c of SOME class => [[class]] | NONE => (case get_type_of_constr_or_abstr thy c of SOME (tyco, _) => (map snd o fst o the) (AList.lookup (op =) ((snd o fst o get_type thy) tyco) c) | NONE => replicate (length vs) []); val the_sort = the o AList.lookup (op =) (map fst vs ~~ sortargs); val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty val chead = build_head thy (c, ty); in Thm.weaken chead Drule.dummy_thm end; fun nothing_cert ctxt c = Nothing (dummy_thm ctxt c); fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, []) | cert_of_eqns ctxt c raw_eqns = let val thy = Proof_Context.theory_of ctxt; val eqns = burrow_fst (canonize_thms thy) raw_eqns; val _ = map (assert_eqn thy) eqns; val (thms, propers) = split_list eqns; val _ = map (fn thm => if c = const_eqn thy thm then () else error ("Wrong head of code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms; val tvars_of = build_rev o Term.add_tvarsT; val vss = map (tvars_of o snd o head_eqn) thms; val inter_sorts = build o fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd); val sorts = map_transpose inter_sorts vss; val vts = Name.invent_names Name.context Name.aT sorts; fun instantiate vs = Thm.instantiate (TVars.make (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts), Vars.empty); val thms' = map2 instantiate vss thms; val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms')))); fun head_conv ct = if can Thm.dest_comb ct then Conv.fun_conv head_conv ct else Conv.rewr_conv head_thm ct; val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv); val cert_thm = Conjunction.intr_balanced (map rewrite_head thms'); in Equations (cert_thm, propers) end; fun cert_of_proj ctxt proj tyco = let val thy = Proof_Context.theory_of ctxt val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco; val _ = if proj = proj' then () else error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj); in Projection (mk_proj tyco vs ty abs proj, tyco) end; fun cert_of_abs ctxt tyco c raw_abs_thm = let val thy = Proof_Context.theory_of ctxt; val abs_thm = singleton (canonize_thms thy) raw_abs_thm; val _ = assert_abs_eqn thy (SOME tyco) abs_thm; val _ = if c = const_abs_eqn thy abs_thm then () else error ("Wrong head of abstract code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy abs_thm); in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; fun constrain_cert_thm thy sorts cert_thm = let val ((vs, _), head) = get_head thy cert_thm; val (subst, cert_thm') = cert_thm |> Thm.implies_intr head |> constrain_thm thy vs sorts; val head' = Thm.term_of head |> subst |> Thm.global_cterm_of thy; val cert_thm'' = cert_thm' |> Thm.elim_implies (Thm.assume head'); in cert_thm'' end; fun constrain_cert thy sorts (Nothing cert_thm) = Nothing (constrain_cert_thm thy sorts cert_thm) | constrain_cert thy sorts (Equations (cert_thm, propers)) = Equations (constrain_cert_thm thy sorts cert_thm, propers) | constrain_cert _ _ (cert as Projection _) = cert | constrain_cert thy sorts (Abstract (abs_thm, tyco)) = Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); fun conclude_cert (Nothing cert_thm) = Nothing (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context) | conclude_cert (Equations (cert_thm, propers)) = Equations (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context, propers) | conclude_cert (cert as Projection _) = cert | conclude_cert (Abstract (abs_thm, tyco)) = Abstract (Thm.close_derivation \<^here> abs_thm |> Thm.trim_context, tyco); fun typscheme_of_cert thy (Nothing cert_thm) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Equations (cert_thm, _)) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Projection (proj, _)) = typscheme_projection thy proj | typscheme_of_cert thy (Abstract (abs_thm, _)) = typscheme_abs thy abs_thm; fun typargs_deps_of_cert thy (Nothing cert_thm) = let val vs = (fst o fst) (get_head thy cert_thm); in (vs, []) end | typargs_deps_of_cert thy (Equations (cert_thm, propers)) = let val vs = (fst o fst) (get_head thy cert_thm); val equations = if null propers then [] else Thm.prop_of cert_thm |> Logic.dest_conjunction_balanced (length propers); in (vs, build (fold (add_rhss_of_eqn thy) equations)) end | typargs_deps_of_cert thy (Projection (t, _)) = (fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = let val vs = fst (typscheme_abs thy abs_thm); val (_, concrete_thm) = concretify_abs thy tyco abs_thm; in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end; fun equations_of_cert thy (cert as Nothing _) = (typscheme_of_cert thy cert, NONE) | equations_of_cert thy (cert as Equations (cert_thm, propers)) = let val tyscm = typscheme_of_cert thy cert; val thms = if null propers then [] else cert_thm |> Thm.transfer thy |> Local_Defs.expand [snd (get_head thy cert_thm)] |> Thm.varifyT_global |> Conjunction.elim_balanced (length propers); fun abstractions (args, rhs) = (map (pair NONE) args, (NONE, rhs)); in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end | equations_of_cert thy (Projection (t, tyco)) = let val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco; val tyscm = typscheme_projection thy t; val t' = Logic.varify_types_global t; fun abstractions (args, rhs) = (map (pair (SOME abs)) args, (NONE, rhs)); in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end | equations_of_cert thy (Abstract (abs_thm, tyco)) = let val tyscm = typscheme_abs thy abs_thm; val (abs, concrete_thm) = concretify_abs thy tyco (Thm.transfer thy abs_thm); fun abstractions (args, rhs) = (map (pair NONE) args, (SOME abs, rhs)); in (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm, (SOME (Thm.varifyT_global abs_thm), true))]) end; fun pretty_cert _ (Nothing _) = [] | pretty_cert thy (cert as Equations _) = (map_filter (Option.map (Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy)) o fst o snd) o these o snd o equations_of_cert thy) cert | pretty_cert thy (Projection (t, _)) = [Syntax.pretty_term_global thy (Logic.varify_types_global t)] | pretty_cert thy (Abstract (abs_thm, _)) = [(Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm]; end; (* code certificate access with preprocessing *) fun eqn_conv conv ct = let fun lhs_conv ct = if can Thm.dest_comb ct then Conv.combination_conv lhs_conv conv ct else Conv.all_conv ct; in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end; fun rewrite_eqn conv ctxt = singleton (Variable.trade (K (map (Conv.fconv_rule (conv (Simplifier.rewrite ctxt))))) ctxt) fun apply_functrans ctxt functrans = let fun trace_eqns s eqns = (Pretty.writeln o Pretty.chunks) (Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns); val tracing = if Config.get ctxt simp_trace then trace_eqns else (K o K) (); in tap (tracing "before function transformation") #> (perhaps o perhaps_loop o perhaps_apply) functrans #> tap (tracing "after function transformation") end; fun preprocess conv ctxt = rewrite_eqn conv ctxt #> Axclass.unoverload ctxt; fun get_cert ctxt functrans c = case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of NONE => nothing_cert ctxt c | SOME (Eqns (_, eqns)) => eqns |> (map o apfst) (Thm.transfer' ctxt) |> apply_functrans ctxt functrans |> (map o apfst) (preprocess eqn_conv ctxt) |> cert_of_eqns ctxt c | SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco | SOME (Abstr (abs_thm, (tyco, _))) => abs_thm |> Thm.transfer' ctxt |> preprocess Conv.arg_conv ctxt |> cert_of_abs ctxt tyco c; (* case certificates *) local fun raw_case_cert thm = let val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; val _ = case head of Free _ => () | Var _ => () | _ => raise TERM ("case_cert", []); val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; val (Const (case_const, _), raw_params) = strip_comb case_expr; val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; val _ = if n = ~1 then raise TERM ("case_cert", []) else (); val params = map (fst o dest_Var) (nth_drop n raw_params); fun dest_case t = let val (head' $ t_co, rhs) = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val (Const (co, _), args) = strip_comb t_co; val (Var (param, _), args') = strip_comb rhs; val _ = if args' = args then () else raise TERM ("case_cert", []); in (param, co) end; fun analyze_cases cases = let val co_list = build (fold (AList.update (op =) o dest_case) cases); in map (AList.lookup (op =) co_list) params end; fun analyze_let t = let val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val _ = if arg' = arg then () else raise TERM ("case_cert", []); val _ = if [param'] = params then () else raise TERM ("case_cert", []); in [] end; fun analyze (cases as [let_case]) = (analyze_cases cases handle Bind => analyze_let let_case) | analyze cases = analyze_cases cases; in (case_const, (n, analyze cases)) end; in fun case_cert thm = raw_case_cert thm handle Bind => error "bad case certificate" | TERM _ => error "bad case certificate"; end; fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy); fun get_case_schema thy c = case lookup_case_spec thy c of SOME (Case {schema, ...}) => SOME schema | _ => NONE; fun get_case_cong thy c = case lookup_case_spec thy c of SOME (Case {cong, ...}) => SOME cong | _ => NONE; fun is_undefined thy c = case lookup_case_spec thy c of SOME Undefined => true | _ => false; (* diagnostic *) fun print_codesetup thy = let val ctxt = Proof_Context.init_global thy; val specs = specs_of thy; fun pretty_equations const thms = (Pretty.block o Pretty.fbreaks) (Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms); fun pretty_function (const, Eqns (_, eqns)) = pretty_equations const (map fst eqns) | pretty_function (const, Proj (proj, _)) = Pretty.block [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; fun pretty_typ (tyco, vs) = Pretty.str (string_of_typ thy (Type (tyco, map TFree vs))); fun pretty_type_spec (typ, (cos, abstract)) = if null cos then pretty_typ typ else (Pretty.block o Pretty.breaks) ( pretty_typ typ :: Pretty.str "=" :: (if abstract then [Pretty.str "(abstract)"] else []) @ separate (Pretty.str "|") (map (fn (c, (_, [])) => Pretty.str (string_of_const thy c) | (c, (_, tys)) => (Pretty.block o Pretty.breaks) (Pretty.str (string_of_const thy c) :: Pretty.str "of" :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) ); fun pretty_case_param NONE = "" | pretty_case_param (SOME (c, _)) = string_of_const thy c fun pretty_case (const, Case {schema = (_, (_, [])), ...}) = Pretty.str (string_of_const thy const) | pretty_case (const, Case {schema = (_, (_, cos)), ...}) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str "with", (Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos] | pretty_case (const, Undefined) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str ""]; val functions = all_fun_specs specs |> sort (string_ord o apply2 fst); val types = History.all (types_of specs) |> map (fn (tyco, (vs, spec)) => ((tyco, vs), constructors_of spec)) |> sort (string_ord o apply2 (fst o fst)); val cases = History.all (cases_of specs) |> filter (fn (_, No_Case) => false | _ => true) |> sort (string_ord o apply2 fst); in Pretty.writeln_chunks [ Pretty.block ( Pretty.str "types:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_type_spec) types ), Pretty.block ( Pretty.str "functions:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_function) functions ), Pretty.block ( Pretty.str "cases:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_case) cases ) ] end; (** declaration of executable ingredients **) (* plugins for dependent applications *) structure Codetype_Plugin = Plugin(type T = string); val codetype_plugin = Plugin_Name.declare_setup \<^binding>\codetype\; fun type_interpretation f = Codetype_Plugin.interpretation codetype_plugin (fn tyco => Local_Theory.background_theory (fn thy => thy |> Sign.root_path |> Sign.add_path (Long_Name.qualifier tyco) |> f tyco |> Sign.restore_naming thy)); fun datatype_interpretation f = type_interpretation (fn tyco => fn thy => case get_type thy tyco of (spec, false) => f (tyco, spec) thy | (_, true) => thy ); fun abstype_interpretation f = type_interpretation (fn tyco => fn thy => case try (get_abstype_spec thy) tyco of SOME spec => f (tyco, spec) thy | NONE => thy ); fun register_tyco_for_plugin tyco = Named_Target.theory_map (Codetype_Plugin.data_default tyco); (* abstract code declarations *) fun code_declaration strictness lift_phi f x = - Local_Theory.declaration {syntax = false, pervasive = false} + Local_Theory.declaration {syntax = false, pervasive = false, pos = Position.thread_data ()} (fn phi => Context.mapping (f strictness (lift_phi phi x)) I); (* types *) fun invalidate_constructors_of (_, type_spec) = fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec)); fun invalidate_abstract_functions_of (_, type_spec) = fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec); fun invalidate_case_combinators_of (_, type_spec) = fold (fn c => History.register c No_Case) (case_combinators_of type_spec); fun register_type (tyco, vs_typ_spec) specs = let val olds = the_list (History.lookup (types_of specs) tyco); in specs |> map_functions (fold invalidate_abstract_functions_of olds #> invalidate_constructors_of vs_typ_spec) |> map_cases (fold invalidate_case_combinators_of olds) |> map_types (History.register tyco vs_typ_spec) end; fun declare_datatype_global proto_constrs thy = let fun unoverload_const_typ (c, ty) = (Axclass.unoverload_const thy (c, ty), ty); val constrs = map unoverload_const_typ proto_constrs; val (tyco, (vs, cos)) = constrset_of_consts thy constrs; in thy |> modify_specs (register_type (tyco, (vs, Constructors {constructors = cos, case_combinators = []}))) |> register_tyco_for_plugin tyco end; fun declare_datatype_cmd raw_constrs thy = declare_datatype_global (map (read_bare_const thy) raw_constrs) thy; fun generic_declare_abstype strictness proto_thm thy = case check_abstype_cert strictness thy proto_thm of SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) => thy |> modify_specs (register_type (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, abs_rep = Thm.trim_context abs_rep, more_abstract_functions = []})) #> register_fun_spec proj (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs)))) |> register_tyco_for_plugin tyco | NONE => thy; val declare_abstype_global = generic_declare_abstype Strict; val declare_abstype = code_declaration Liberal Morphism.thm generic_declare_abstype; (* functions *) (* strictness wrt. shape of theorem propositions: * default equations: silent * using declarations and attributes: warnings (after morphism application!) * using global declarations (... -> thy -> thy): strict * internal processing after storage: strict *) local fun subsumptive_add thy verbose (thm, proper) eqns = let val args_of = drop_prefix is_Var o rev o snd o strip_comb o Term.map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of o Thm.transfer thy; val args = args_of thm; val incr_idx = Logic.incr_indexes ([], [], Thm.maxidx_of thm + 1); fun matches_args args' = let val k = length args' - length args in if k >= 0 then Pattern.matchess thy (args, (map incr_idx o drop k) args') else false end; fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then (if verbose then warning ("Code generator: dropping subsumed code equation\n" ^ Thm.string_of_thm_global thy thm') else (); true) else false; in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end; fun add_eqn_for (c, eqn) thy = thy |> modify_specs (modify_pending_eqns c (subsumptive_add thy true eqn)); fun add_eqns_for default (c, proto_eqns) thy = thy |> modify_specs (fn specs => if is_default (lookup_fun_spec specs c) orelse not default then let val eqns = [] |> fold_rev (subsumptive_add thy (not default)) proto_eqns; in specs |> register_fun_spec c (Eqns (default, eqns)) end else specs); fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) = modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm |> Thm.trim_context, tyco_abs)) #> map_types (History.modify_entry tyco (add_abstract_function c))) in fun generic_declare_eqns default strictness raw_eqns thy = fold (add_eqns_for default) (prep_eqns strictness thy raw_eqns) thy; fun generic_add_eqn strictness raw_eqn thy = fold add_eqn_for (the_list (prep_eqn strictness thy raw_eqn)) thy; fun generic_declare_abstract_eqn strictness raw_abs_eqn thy = fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy; fun add_maybe_abs_eqn_liberal thm thy = case prep_maybe_abs_eqn thy thm of SOME (c, (eqn, NONE)) => add_eqn_for (c, eqn) thy | SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy | NONE => thy; end; val declare_default_eqns_global = generic_declare_eqns true Silent; val declare_default_eqns = code_declaration Silent (map o apfst o Morphism.thm) (generic_declare_eqns true); val declare_eqns_global = generic_declare_eqns false Strict; val declare_eqns = code_declaration Liberal (map o apfst o Morphism.thm) (generic_declare_eqns false); val add_eqn_global = generic_add_eqn Strict; fun del_eqn_global thm thy = case prep_eqn Liberal thy (thm, false) of SOME (c, (thm, _)) => modify_specs (modify_pending_eqns c (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')))) thy | NONE => thy; val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict; val declare_abstract_eqn = code_declaration Liberal Morphism.thm generic_declare_abstract_eqn; fun declare_aborting_global c = modify_specs (register_fun_spec c aborting); fun declare_unimplemented_global c = modify_specs (register_fun_spec c unimplemented); (* cases *) fun case_cong thy case_const (num_args, (pos, _)) = let val ([x, y], ctxt) = fold_map Name.variant ["A", "A'"] Name.context; val (zs, _) = fold_map Name.variant (replicate (num_args - 1) "") ctxt; val (ws, vs) = chop pos zs; val T = devarify (const_typ thy case_const); val Ts = binder_types T; val T_cong = nth Ts pos; fun mk_prem z = Free (z, T_cong); fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts); val (prem, concl) = apply2 Logic.mk_equals (apply2 mk_prem (x, y), apply2 mk_concl (x, y)); in Goal.prove_sorry_global thy (x :: y :: zs) [prem] concl (fn {context = ctxt', prems} => Simplifier.rewrite_goals_tac ctxt' prems THEN ALLGOALS (Proof_Context.fact_tac ctxt' [Drule.reflexive_thm])) end; fun declare_case_global thm thy = let val (case_const, (k, cos)) = case_cert thm; fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of SOME (c, false) => SOME c | _ => NONE; val cos_with_tycos = (map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos; val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of [] => () | cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs); val tycos = distinct (op =) (map_filter snd cos_with_tycos); val schema = (1 + Int.max (1, length cos), (k, (map o Option.map) (fn c => (c, args_number thy c)) cos)); val cong = case_cong thy case_const schema; in thy |> modify_specs (map_cases (History.register case_const (Case {schema = schema, tycos = tycos, cong = cong})) #> map_types (fold (fn tyco => History.modify_entry tyco (add_case_combinator case_const)) tycos)) end; fun declare_undefined_global c = (modify_specs o map_cases) (History.register c Undefined); (* attributes *) fun code_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); fun code_thm_attribute g f = Scan.lift (g |-- Scan.succeed (code_attribute f)); fun code_const_attribute g f = Scan.lift (g -- Args.colon) |-- Scan.repeat1 Args.term >> (fn ts => code_attribute (K (fold (fn t => fn thy => f ((check_const thy o Logic.unvarify_types_global) t) thy) ts))); val _ = Theory.setup (let val code_attribute_parser = code_thm_attribute (Args.$$$ "equation") (fn thm => generic_add_eqn Liberal (thm, true)) || code_thm_attribute (Args.$$$ "nbe") (fn thm => generic_add_eqn Liberal (thm, false)) || code_thm_attribute (Args.$$$ "abstract") (generic_declare_abstract_eqn Liberal) || code_thm_attribute (Args.$$$ "abstype") (generic_declare_abstype Liberal) || code_thm_attribute Args.del del_eqn_global || code_const_attribute (Args.$$$ "abort") declare_aborting_global || code_const_attribute (Args.$$$ "drop") declare_unimplemented_global || Scan.succeed (code_attribute add_maybe_abs_eqn_liberal); in Attrib.setup \<^binding>\code\ code_attribute_parser "declare theorems for code generation" end); end; (*struct*) (* type-safe interfaces for data dependent on executable code *) functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = struct type T = Data.T; exception Data of T; fun dest (Data x) = x val kind = Code.declare_data (Data Data.empty); val data_op = (kind, Data, dest); fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f | change_yield NONE f = f Data.empty fun change some_thy f = snd (change_yield some_thy (pair () o f)); end; structure Code : CODE = struct open Code; end; diff --git a/src/Pure/Isar/entity.ML b/src/Pure/Isar/entity.ML --- a/src/Pure/Isar/entity.ML +++ b/src/Pure/Isar/entity.ML @@ -1,58 +1,59 @@ (* Title: Pure/Isar/entity.ML Author: Makarius Entity definitions within a global or local theory context. *) signature ENTITY = sig type 'a data_ops = {get_data: Context.generic -> 'a Name_Space.table, put_data: 'a Name_Space.table -> Context.generic -> Context.generic} val define_global: 'a data_ops -> binding -> 'a -> theory -> string * theory val define: 'a data_ops -> binding -> 'a -> local_theory -> string * local_theory end; structure Entity: ENTITY = struct (* context data *) type 'a data_ops = {get_data: Context.generic -> 'a Name_Space.table, put_data: 'a Name_Space.table -> Context.generic -> Context.generic}; (* global definition (foundation) *) fun define_global {get_data, put_data} b x thy = let val context = Context.Theory thy; val (name, data') = Name_Space.define context true (b, x) (get_data context); in (name, Context.the_theory (put_data data' context)) end; (* local definition *) fun alias {get_data, put_data} binding name = - Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val naming = Name_Space.naming_of context; - val binding' = Morphism.binding phi binding; - val data' = Name_Space.alias_table naming binding' name (get_data context); - in put_data data' context end); + Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} + (fn phi => fn context => + let + val naming = Name_Space.naming_of context; + val binding' = Morphism.binding phi binding; + val data' = Name_Space.alias_table naming binding' name (get_data context); + in put_data data' context end); fun transfer {get_data, put_data} ctxt = let val data0 = get_data (Context.Theory (Proof_Context.theory_of ctxt)); val data' = Name_Space.merge_tables (data0, get_data (Context.Proof ctxt)); in Context.proof_map (put_data data') ctxt end; fun define ops binding x = Local_Theory.background_theory_result (define_global ops binding x) #-> (fn name => Local_Theory.map_contexts (K (transfer ops)) #> alias ops binding name #> pair name); end; diff --git a/src/Pure/Isar/expression.ML b/src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML +++ b/src/Pure/Isar/expression.ML @@ -1,881 +1,881 @@ (* Title: Pure/Isar/expression.ML Author: Clemens Ballarin, TU Muenchen Locale expressions and user interface layer of locales. *) signature EXPRESSION = sig (* Locale expressions *) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list type 'term rewrites = (Attrib.binding * 'term) list type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list type expression_i = (string, term) expr * (binding * typ option * mixfix) list type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list (* Processing of context statements *) val cert_statement: Element.context_i list -> Element.statement_i -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context val read_statement: Element.context list -> Element.statement -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context (* Declaring locales *) val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) (*FIXME*) val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val add_locale: binding -> binding -> Bundle.name list -> expression_i -> Element.context_i list -> theory -> string * local_theory val add_locale_cmd: binding -> binding -> (xstring * Position.T) list -> expression -> Element.context list -> theory -> string * local_theory (* Processing of locale expressions *) val cert_goal_expression: expression_i -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context val read_goal_expression: expression -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context end; structure Expression : EXPRESSION = struct datatype ctxt = datatype Element.ctxt; (*** Expressions ***) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list; type 'term rewrites = (Attrib.binding * 'term) list; type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list; type expression_i = (string, term) expr * (binding * typ option * mixfix) list; type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; (** Internalise locale names in expr **) fun check_expr thy instances = map (apfst (Locale.check thy)) instances; (** Parameters of expression **) (*Sanity check of instantiations and extraction of implicit parameters. The latter only occurs iff strict = false. Positional instantiations are extended to match full length of parameter list of instantiated locale.*) fun parameters_of thy strict (expr, fixed) = let val ctxt = Proof_Context.init_global thy; fun reject_dups message xs = (case duplicates (op =) xs of [] => () | dups => error (message ^ commas dups)); fun parm_eq ((p1, mx1), (p2, mx2)) = p1 = p2 andalso (Mixfix.equal (mx1, mx2) orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^ Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); fun params_inst (loc, (prfx, (Positional insts, eqns))) = let val ps = params_loc loc; val d = length ps - length insts; val insts' = if d < 0 then error ("More arguments than parameters in instantiation of locale " ^ quote (Locale.markup_name ctxt loc)) else insts @ replicate d NONE; val ps' = (ps ~~ insts') |> map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); in (ps', (loc, (prfx, (Positional insts', eqns)))) end | params_inst (loc, (prfx, (Named insts, eqns))) = let val _ = reject_dups "Duplicate instantiation of the following parameter(s): " (map fst insts); val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => if AList.defined (op =) ps p then AList.delete (op =) p ps else error (quote p ^ " not a parameter of instantiated expression")); in (ps', (loc, (prfx, (Named insts, eqns)))) end; fun params_expr is = let val (is', ps') = fold_map (fn i => fn ps => let val (ps', i') = params_inst i; val ps'' = distinct parm_eq (ps @ ps'); in (i', ps'') end) is [] in (ps', is') end; val (implicit, expr') = params_expr expr; val implicit' = map #1 implicit; val fixed' = map (Variable.check_name o #1) fixed; val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; val implicit'' = if strict then [] else let val _ = reject_dups "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed'); in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; in (expr', implicit'' @ fixed) end; (** Read instantiation **) (* Parse positional or named instantiation *) local fun prep_inst prep_term ctxt parms (Positional insts) = (insts ~~ parms) |> map (fn (NONE, p) => Free (p, dummyT) | (SOME t, _) => prep_term ctxt t) | prep_inst prep_term ctxt parms (Named insts) = parms |> map (fn p => (case AList.lookup (op =) insts p of SOME t => prep_term ctxt t | NONE => Free (p, dummyT))); in fun parse_inst x = prep_inst Syntax.parse_term x; fun make_inst x = prep_inst (K I) x; end; (* Instantiation morphism *) fun inst_morphism params ((prfx, mandatory), insts') ctxt = let (* parameters *) val parm_types = map #2 params; val type_parms = fold Term.add_tfreesT parm_types []; (* type inference *) val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; val type_parms' = fold Term.add_tvarsT parm_types' []; val checked = (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) val (type_parms'', insts'') = chop (length type_parms') checked; (* context *) val ctxt' = fold Proof_Context.augment checked ctxt; val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; (* instantiation *) val instT = TFrees.build (fold2 (fn v => fn T => not (TFree v = T) ? TFrees.add (v, T)) type_parms (map Logic.dest_type type_parms'')); val cert_inst = Frees.build (fold2 (fn v => fn t => not (Free v = t) ? Frees.add (v, cert t)) (map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) insts''); in (Element.instantiate_normalize_morphism (TFrees.map (K certT) instT, cert_inst) $> Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') end; (*** Locale processing ***) (** Parsing **) fun parse_elem prep_typ prep_term ctxt = Element.map_ctxt {binding = I, typ = prep_typ ctxt, term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), fact = I, attrib = I}; fun prepare_stmt prep_prop prep_obtains ctxt stmt = (case stmt of Element.Shows raw_shows => raw_shows |> (map o apsnd o map) (fn (t, ps) => (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) | Element.Obtains raw_obtains => let val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); (** Simultaneous type inference: instantiations + elements + statement **) local fun mk_type T = (Logic.mk_type T, []); fun mk_term t = (t, []); fun mk_propp (p, pats) = (Type.constraint propT p, pats); fun dest_type (T, []) = Logic.dest_type T; fun dest_term (t, []) = t; fun dest_propp (p, pats) = (p, pats); fun extract_inst (_, (_, ts)) = map mk_term ts; fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)); fun extract_eqns es = map (mk_term o snd) es; fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs; fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | extract_elem (Notes _) = [] | extract_elem (Lazy_Notes _) = []; fun restore_elem (Fixes fixes, css) = (fixes ~~ css) |> map (fn ((x, _, mx), cs) => (x, cs |> map dest_type |> try hd, mx)) |> Fixes | restore_elem (Constrains csts, css) = (csts ~~ css) |> map (fn ((x, _), cs) => (x, cs |> map dest_type |> hd)) |> Constrains | restore_elem (Assumes asms, css) = (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes | restore_elem (Defines defs, css) = (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | restore_elem (elem as Notes _, _) = elem | restore_elem (elem as Lazy_Notes _, _) = elem; fun prep (_, pats) (ctxt, t :: ts) = let val ctxt' = Proof_Context.augment t ctxt in ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), (ctxt', ts)) end; fun check cs ctxt = let val (cs', (ctxt', _)) = fold_map prep cs (ctxt, Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); in (cs', ctxt') end; in fun check_autofix insts eqnss elems concl ctxt = let val inst_cs = map extract_inst insts; val eqns_cs = map extract_eqns eqnss; val elem_css = map extract_elem elems; val concl_cs = (map o map) mk_propp (map snd concl); (* Type inference *) val (inst_cs' :: eqns_cs' :: css', ctxt') = (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; val (elem_css', [concl_cs']) = chop (length elem_css) css'; in ((map restore_inst (insts ~~ inst_cs'), map restore_eqns (eqnss ~~ eqns_cs'), map restore_elem (elems ~~ elem_css'), map fst concl ~~ concl_cs'), ctxt') end; end; (** Prepare locale elements **) fun declare_elem prep_var (Fixes fixes) ctxt = let val (vars, _) = fold_map prep_var fixes ctxt in ctxt |> Proof_Context.add_fixes vars |> snd end | declare_elem prep_var (Constrains csts) ctxt = ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd | declare_elem _ (Assumes _) ctxt = ctxt | declare_elem _ (Defines _) ctxt = ctxt | declare_elem _ (Notes _) ctxt = ctxt | declare_elem _ (Lazy_Notes _) ctxt = ctxt; (** Finish locale elements **) fun finish_inst ctxt (loc, (prfx, inst)) = let val thy = Proof_Context.theory_of ctxt; val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; in (loc, morph) end; fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => let val x = Binding.name_of binding in (binding, AList.lookup (op =) parms x, mx) end); local fun closeup _ _ false elem = elem | closeup (outer_ctxt, ctxt) parms true elem = let (* FIXME consider closing in syntactic phase -- before type checking *) fun close_frees t = let val rev_frees = Term.fold_aterms (fn Free (x, T) => if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t []; in fold (Logic.all o Free) rev_frees t end; fun no_binds [] = [] | no_binds _ = error "Illegal term bindings in context element"; in (case elem of Assumes asms => Assumes (asms |> map (fn (a, propps) => (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | e => e) end; in fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | finish_elem _ _ _ (Constrains _) = Constrains [] | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | finish_elem _ _ _ (elem as Notes _) = elem | finish_elem _ _ _ (elem as Lazy_Notes _) = elem; end; (** Process full context statement: instantiations + elements + statement **) (* Interleave incremental parsing and type inference over entire parsed stretch. *) local fun abs_def ctxt = Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; fun prep_full_context_statement parse_typ parse_prop prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt ctxt1 = let val thy = Proof_Context.theory_of ctxt1; val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = let val params = map #1 (Locale.params_of thy loc); val inst' = prep_inst ctxt (map #1 params) inst; val parm_types' = params |> map (#2 #> Logic.varifyT_global #> Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> Type_Infer.paramify_vars); val inst'' = map2 Type.constraint parm_types' inst'; val insts' = insts @ [(loc, (prfx, inst''))]; val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt; val inst''' = insts'' |> List.last |> snd |> snd; val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt; val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 handle ERROR msg => if null eqns then error msg else (Locale.tracing ctxt1 (msg ^ "\nFalling back to reading rewrites clause before activation."); ctxt2); val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns; val eqns' = (prep_eqns ctxt' o map snd) eqns; val eqnss' = [attrss ~~ eqns']; val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt'; val rewrite_morph = eqns' |> map (abs_def ctxt') |> Variable.export_terms ctxt' ctxt |> Element.eq_term_morphism |> Morphism.default; val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt; val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns']; in (i + 1, insts', eqnss', ctxt'') end; fun prep_elem raw_elem ctxt = let val ctxt' = ctxt |> Context_Position.set_visible false |> declare_elem prep_var_elem raw_elem |> Context_Position.restore_visible ctxt; val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem; in (elems', ctxt') end; val fors = fold_map prep_var_inst fixed ctxt1 |> fst; val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2); fun prep_stmt elems ctxt = check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt; val _ = if fixed_frees then () else (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of [] => () | frees => error ("Illegal free variables in expression: " ^ commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))); val ((insts, _, elems', concl), ctxt4) = ctxt3 |> init_body |> fold_map prep_elem raw_elems |-> prep_stmt; (* parameters from expression and elements *) val xs = maps (fn Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) (Fixes fors :: elems'); val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4; val fors' = finish_fixes parms fors; val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; val deps = map (finish_inst ctxt5) insts; val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems'; in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end; in fun cert_full_context_statement x = prep_full_context_statement (K I) (K I) Obtain.cert_obtains Proof_Context.cert_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun cert_read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src Proof_Context.read_var check_expr x; end; (* Context statement: elements + statement *) local fun prep_statement prep activate raw_elems raw_stmt ctxt = let val ((_, _, _, elems, concl), _) = prep {strict = true, do_close = false, fixed_frees = true} ([], []) I raw_elems raw_stmt ctxt; val ctxt' = ctxt |> Proof_Context.set_stmt true |> fold_map activate elems |> #2 |> Proof_Context.restore_stmt ctxt; in (concl, ctxt') end; in fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x; fun read_statement x = prep_statement read_full_context_statement Element.activate x; end; (* Locale declaration: import + elements *) fun fix_params params = Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd; local fun prep_declaration prep activate raw_import init_body raw_elems ctxt = let val ((fixed, deps, eqnss, elems, _), (parms, ctxt0)) = prep {strict = false, do_close = true, fixed_frees = false} raw_import init_body raw_elems (Element.Shows []) ctxt; val _ = null (flat eqnss) orelse error "Illegal rewrites clause(s) in declaration of locale"; (* Declare parameters and imported facts *) val ctxt' = ctxt |> fix_params fixed |> fold (Context.proof_map o Locale.activate_facts NONE) deps; val (elems', ctxt'') = ctxt' |> Proof_Context.set_stmt true |> fold_map activate elems ||> Proof_Context.restore_stmt ctxt'; in ((fixed, deps, elems', ctxt''), (parms, ctxt0)) end; in fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x; fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x; fun read_declaration x = prep_declaration read_full_context_statement Element.activate x; end; (* Locale expression to set up a goal *) local fun props_of thy (name, morph) = let val (asm, defs) = Locale.specification_of thy name in map (Morphism.term morph) (the_list asm @ defs) end; fun prep_goal_expression prep expression ctxt = let val thy = Proof_Context.theory_of ctxt; val ((fixed, deps, eqnss, _, _), _) = prep {strict = true, do_close = true, fixed_frees = true} expression I [] (Element.Shows []) ctxt; (* proof obligations *) val propss = map (props_of thy) deps; val eq_propss = (map o map) snd eqnss; val goal_ctxt = ctxt |> fix_params fixed |> (fold o fold) Proof_Context.augment (propss @ eq_propss); val export = Proof_Context.export_morphism goal_ctxt ctxt; val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; val exp_term = Term_Subst.zero_var_indexes o Morphism.term export; val exp_typ = Logic.type_map exp_term; val export' = Morphism.morphism "Expression.prep_goal" {binding = [], typ = [K exp_typ], term = [K exp_term], fact = [K exp_fact]}; in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end; in fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x; fun read_goal_expression x = prep_goal_expression read_full_context_statement x; end; (*** Locale declarations ***) (* extract specification text *) val norm_term = Envir.beta_norm oo Term.subst_atomic; fun bind_def ctxt eq (env, eqs) = let val _ = Local_Defs.cert_def ctxt (K []) eq; val ((y, T), b) = Local_Defs.abs_def eq; val b' = norm_term env b; fun err msg = error (msg ^ ": " ^ quote y); in (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of [] => ((Free (y, T), b') :: env, eq :: eqs) | dups => if forall (fn (_, b'') => b' aconv b'') dups then (env, eqs) else err "Attempt to redefine variable") end; (* text has the following structure: (((exts, exts'), (ints, ints')), (xs, env, defs)) where exts: external assumptions (terms in assumes elements) exts': dito, normalised wrt. env ints: internal assumptions (terms in assumptions from insts) ints': dito, normalised wrt. env xs: the free variables in exts' and ints' and rhss of definitions, this includes parameters except defined parameters env: list of term pairs encoding substitutions, where the first term is a free variable; substitutions represent defines elements and the rhs is normalised wrt. the previous env defs: the equations from the defines elements *) fun eval_text _ _ (Fixes _) text = text | eval_text _ _ (Constrains _) text = text | eval_text _ is_ext (Assumes asms) (((exts, exts'), (ints, ints')), (env, defs)) = let val ts = maps (map #1 o #2) asms; val ts' = map (norm_term env) ts; val spec' = if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) else ((exts, exts'), (ints @ ts, ints' @ ts')); in (spec', (env, defs)) end | eval_text ctxt _ (Defines defs) (spec, binds) = (spec, fold (bind_def ctxt o #1 o #2) defs binds) | eval_text _ _ (Notes _) text = text | eval_text _ _ (Lazy_Notes _) text = text; fun eval_inst ctxt (loc, morph) text = let val thy = Proof_Context.theory_of ctxt; val (asm, defs) = Locale.specification_of thy loc; val asm' = Option.map (Morphism.term morph) asm; val defs' = map (Morphism.term morph) defs; val text' = text |> (if is_some asm then eval_text ctxt false (Assumes [(Binding.empty_atts, [(the asm', [])])]) else I) |> (if not (null defs) then eval_text ctxt false (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs')) else I) (* FIXME clone from locale.ML *) in text' end; fun eval_elem ctxt elem text = eval_text ctxt true elem text; fun eval ctxt deps elems = let val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [])); val ((spec, (_, defs))) = fold (eval_elem ctxt) elems text'; in (spec, defs) end; (* axiomsN: name of theorem set with destruct rules for locale predicates, also name suffix of delta predicates and assumptions. *) val axiomsN = "axioms"; local (* introN: name of theorems for introduction rules of locale and delta predicates *) val introN = "intro"; fun atomize_spec ctxt ts = let val t = Logic.mk_conjunction_balanced ts; val body = Object_Logic.atomize_term ctxt t; val bodyT = Term.fastype_of body; in if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of ctxt t)) else (body, bodyT, Object_Logic.atomize ctxt (Thm.cterm_of ctxt t)) end; (* achieve plain syntax for locale predicates (without "PROP") *) fun aprop_tr' n c = let val c' = Lexicon.mark_const c; fun tr' (_: Proof.context) T args = if T <> dummyT andalso length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.const c', args) else raise Match; in (c', tr') end; (* define one predicate including its intro rule and axioms - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) - norm_ts: terms representing locale assumptions (normalised wrt. defs) - thy: the theory *) fun def_pred binding parms defs ts norm_ts thy = let val name = Sign.full_name thy binding; val thy_ctxt = Proof_Context.init_global thy; val (body, bodyT, body_eq) = atomize_spec thy_ctxt norm_ts; val env = Names.build (Names.add_free_names body); val xs = filter (Names.defined env o #1) parms; val Ts = map #2 xs; val type_tfrees = TFrees.build (fold TFrees.add_tfreesT Ts); val extra_tfrees = TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) body) |> TFrees.keys |> map TFree; val predT = map Term.itselfT extra_tfrees ---> Ts ---> bodyT; val args = map Logic.mk_type extra_tfrees @ map Free xs; val head = Term.list_comb (Const (name, predT), args); val statement = Object_Logic.ensure_propT thy_ctxt head; val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.typed_print_translation [aprop_tr' (length args) name] |> Sign.declare_const_global ((binding, predT), NoSyn) |> snd |> Global_Theory.add_defs false [((Thm.def_binding binding, Logic.mk_equals (head, body)), [])]; val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; val intro = Goal.prove_global defs_thy [] norm_ts statement (fn {context = ctxt, ...} => rewrite_goals_tac ctxt [pred_def] THEN compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN compose_tac defs_ctxt (false, Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1); val conjuncts = (Drule.equal_elim_rule2 OF [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_ctxt statement))]) |> Conjunction.elim_balanced (length ts); val (_, axioms_ctxt) = defs_ctxt |> Assumption.add_assumes (maps Thm.chyps_of (defs @ conjuncts)); val axioms = ts ~~ conjuncts |> map (fn (t, ax) => Element.prove_witness axioms_ctxt t (rewrite_goals_tac axioms_ctxt defs THEN compose_tac axioms_ctxt (false, ax, 0) 1)); in ((statement, intro, axioms), defs_thy) end; in (* main predicate definition function *) fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val ctxt = Proof_Context.init_global thy; val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs; val (a_pred, a_intro, a_axioms, thy2) = if null exts then (NONE, NONE, [], thy) else let val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy1) = thy |> def_pred abinding parms defs' exts exts'; val ((_, [intro']), thy2) = thy1 |> Sign.qualified_path true abinding |> Global_Theory.note_thms "" ((Binding.name introN, []), [([intro], [Locale.unfold_add])]) ||> Sign.restore_naming thy1; in (SOME statement, SOME intro', axioms, thy2) end; val (b_pred, b_intro, b_axioms, thy4) = if null ints then (NONE, NONE, [], thy2) else let val ((statement, intro, axioms), thy3) = thy2 |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val conclude_witness = Drule.export_without_context o Element.conclude_witness (Proof_Context.init_global thy3); val ([(_, [intro']), _], thy4) = thy3 |> Sign.qualified_path true binding |> Global_Theory.note_thmss "" [((Binding.name introN, []), [([intro], [Locale.intro_add])]), ((Binding.name axiomsN, []), [(map conclude_witness axioms, [])])] ||> Sign.restore_naming thy3; in (SOME statement, SOME intro', axioms, thy4) end; in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy4) end; end; local fun assumes_to_notes (Assumes asms) axms = fold_map (fn (a, spec) => fn axs => let val (ps, qs) = chop (length spec) axs in ((a, [(ps, [])]), qs) end) asms axms |> apfst (curry Notes "") | assumes_to_notes e axms = (e, axms); fun defines_to_notes ctxt (Defines defs) = Notes ("", map (fn (a, (def, _)) => (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], - [Attrib.internal (K Locale.witness_add)])])) defs) + [Attrib.internal @{here} (K Locale.witness_add)])])) defs) | defines_to_notes _ e = e; val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false; fun gen_add_locale prep_include prep_decl binding raw_predicate_binding raw_includes raw_import raw_body thy = let val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); val ctxt = Proof_Context.init_global thy; val includes = map (prep_include ctxt) raw_includes; val ((fixed, deps, body_elems, _), (parms, ctxt')) = ctxt |> Bundle.includes includes |> prep_decl raw_import I raw_body; val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; val type_tfrees = TFrees.build (fold (TFrees.add_tfreesT o #2) parms); val extra_tfrees = TFrees.build (fold (TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) exts') |> TFrees.keys; val _ = if null extra_tfrees then () else warning ("Additional type variable(s) in locale specification " ^ Binding.print binding ^ ": " ^ commas (map (Syntax.string_of_typ ctxt' o TFree) extra_tfrees)); val predicate_binding = if Binding.is_empty raw_predicate_binding then binding else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = define_preds predicate_binding parms text thy; val pred_ctxt = Proof_Context.init_global thy'; val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; val params = fixed @ maps (fn Fixes fixes => map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems; val asm = if is_some b_statement then b_statement else a_statement; val hyp_spec = filter is_hyp body_elems; val notes = if is_some asm then [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []), [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))], - [Attrib.internal (K Locale.witness_add)])])])] + [Attrib.internal @{here} (K Locale.witness_add)])])])] else []; val notes' = body_elems |> map (Element.transfer_ctxt thy') |> map (defines_to_notes pred_ctxt) |> map (Element.transform_ctxt a_satisfy) |> (fn elems => fold_map assumes_to_notes elems (map (Element.conclude_witness pred_ctxt) a_axioms)) |> fst |> map (Element.transform_ctxt b_satisfy) |> map_filter (fn Notes notes => SOME notes | _ => NONE); val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; val axioms = map (Element.conclude_witness pred_ctxt) b_axioms; val loc_ctxt = thy' |> Locale.register_locale binding (extra_tfrees, params) (asm, rev defs) (a_intro, b_intro) axioms hyp_spec [] (rev notes) (rev deps') |> Named_Target.init includes name |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; in val add_locale = gen_add_locale (K I) cert_declaration; val add_locale_cmd = gen_add_locale Bundle.check read_declaration; end; end; diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,480 +1,481 @@ (* Title: Pure/Isar/generic_target.ML Author: Makarius Author: Florian Haftmann, TU Muenchen Common target infrastructure. *) signature GENERIC_TARGET = sig (*auxiliary*) val export_abbrev: Proof.context -> (term -> term) -> term -> term * ((string * sort) list * (term list * term list)) val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix val check_mixfix_global: binding * bool -> mixfix -> mixfix (*background primitives*) val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val background_declaration: Morphism.declaration_entity -> local_theory -> local_theory val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> theory -> theory (*nested local theories primitives*) val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> local_theory -> local_theory val standard_declaration: (int * int -> bool) -> Morphism.declaration_entity -> local_theory -> local_theory val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val local_interpretation: Locale.registration -> local_theory -> local_theory (*lifting target primitives to local theory operations*) val define: (((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory) -> bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val notes: (string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: (Syntax.mode -> binding * mixfix -> term -> term list * term list -> local_theory -> local_theory) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory (*theory target primitives*) val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory (*theory target operations*) val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val theory_declaration: Morphism.declaration_entity -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory (*locale target primitives*) val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory - val locale_target_declaration: string -> bool -> Morphism.declaration_entity -> - local_theory -> local_theory + val locale_target_declaration: string -> {syntax: bool, pos: Position.T} -> + Morphism.declaration_entity -> local_theory -> local_theory val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory (*locale operations*) val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory - val locale_declaration: string -> {syntax: bool, pervasive: bool} -> Morphism.declaration_entity -> - local_theory -> local_theory + val locale_declaration: string -> {syntax: bool, pervasive: bool, pos: Position.T} -> + Morphism.declaration_entity -> local_theory -> local_theory val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val locale_dependency: string -> Locale.registration -> local_theory -> local_theory end structure Generic_Target: GENERIC_TARGET = struct (** consts **) fun export_abbrev lthy preprocess rhs = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); val rhs' = rhs |> Assumption.export_term lthy (Local_Theory.target_of lthy) |> preprocess; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); val u = fold_rev lambda term_params rhs'; val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; val type_tfrees = TFrees.build (TFrees.add_tfreesT (Term.fastype_of u)); val extra_tfrees = TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) u) |> TFrees.keys; val type_params = map (Logic.mk_type o TFree) extra_tfrees; in (global_rhs, (extra_tfrees, (type_params, term_params))) end; fun check_mixfix ctxt (b, extra_tfrees) mx = if null extra_tfrees then mx else (if Context_Position.is_visible ctxt then warning ("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o TFree) extra_tfrees) ^ (if Mixfix.is_empty mx then "" else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) else (); NoSyn); fun check_mixfix_global (b, no_params) mx = if no_params orelse Mixfix.is_empty mx then mx else (warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^ Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); fun same_const (Const (c, _), Const (c', _)) = c = c' | same_const (t $ _, t' $ _) = same_const (t, t') | same_const (_, _) = false; fun const_decl phi_pred prmode ((b, mx), rhs) = Morphism.entity (fn phi => fn context => if phi_pred phi then let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; val same_shape = Term.aconv_untyped (rhs, rhs'); val same_stem = same_shape orelse same_const (rhs, rhs'); val const_alias = if same_shape then (case rhs' of Const (c, T) => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; in (case Type_Infer_Context.const_type ctxt c of SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE | NONE => NONE) end | _ => NONE) else NONE; in (case const_alias of SOME c => context |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |> Morphism.form_entity (Proof_Context.generic_notation true prmode [(rhs', mx)]) | NONE => context |> Proof_Context.generic_add_abbrev Print_Mode.internal (b', Term.close_schematic_term rhs') |-> (fn (const as Const (c, _), _) => same_stem ? (Proof_Context.generic_revert_abbrev (#1 prmode) c #> same_shape ? Morphism.form_entity (Proof_Context.generic_notation true prmode [(const, mx)])))) end else context); (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = ((binding * (term * term list) -> Context.generic -> Context.generic) * stamp) list val empty = []; val merge = Library.merge (eq_snd (op =)); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (cons (f, stamp ())); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = fold (fn (f, _) => f binding_const_params) interps; in lthy |> Local_Theory.background_theory (Context.theory_map interp) |> Local_Theory.map_contexts (K (Context.proof_map interp)) end; fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = let val params = type_params @ term_params; val target_params = type_params @ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); val mx' = check_mixfix_global (b, null params) mx; val (const, lthy2) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); val lhs = Term.list_comb (const, params); val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result (Thm.add_def (Proof_Context.defs_context lthy2) false false (Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) ||> foundation_interpretation (b, (const, target_params)); in ((lhs, def), lthy3) end; fun background_declaration decl lthy = let fun theory_decl context = Local_Theory.standard_form lthy (Proof_Context.init_global (Context.theory_of context)) decl context; in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; fun background_abbrev (b, global_rhs) params = Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) #>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) (** nested local theories primitives **) fun standard_facts lthy ctxt = Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); fun standard_notes pred kind facts lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd else ctxt) lthy; fun standard_declaration pred decl lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt else ctxt) lthy; fun standard_const pred prmode ((b, mx), rhs) = standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); fun standard_registration pred registration lthy = Local_Theory.map_contexts (fn level => if pred (Local_Theory.level lthy, level) then Context.proof_map (Locale.add_registration registration) else I) lthy; val local_interpretation = standard_registration (fn (n, level) => level >= n - 1); (** lifting target primitives to local theory operations **) (* define *) fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) val ((defs, _), rhs') = Thm.cterm_of lthy rhs |> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; val xs = Variable.add_fixed lthy rhs' []; val T = Term.fastype_of rhs; val type_tfrees = TFrees.build (TFrees.add_tfreesT T #> fold (TFrees.add_tfreesT o #2) xs); val extra_tfrees = TFrees.build (rhs |> TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) |> TFrees.keys; val mx' = check_mixfix lthy (b, extra_tfrees) mx; val type_params = map (Logic.mk_type o TFree) extra_tfrees; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); val params = type_params @ term_params; val U = map Term.fastype_of params ---> T; (*foundation*) val ((lhs', global_def), lthy2) = lthy |> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); (*local definition*) val ([(lhs, (_, local_def))], lthy3) = lthy2 |> Context_Position.set_visible false |> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] ||> Context_Position.restore_visible lthy2; (*result*) val def = Thm.transitive local_def global_def |> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); val ([(res_name, [res])], lthy4) = lthy3 |> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; in ((lhs, (res_name, res)), lthy4) end; (* notes *) local fun import_export_proof ctxt (name, raw_th) = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); (*export assumes/defines*) val th = Goal.norm_result ctxt raw_th; val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; (*export fixes*) val tfrees = TFrees.build (Thm.fold_terms {hyps = true} TFrees.add_tfrees th') |> TFrees.keys |> map TFree; val frees = Frees.build (Thm.fold_terms {hyps = true} Frees.add_frees th') |> Frees.list_set_rev |> map Free; val (th'' :: vs) = (th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; (*thm definition*) val result = Global_Theory.name_thm Global_Theory.official1 name th''; (*import fixes*) val (tvars, vars) = chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |>> map Logic.dest_type; val instT = TVars.build (fold2 (fn a => fn b => (case a of TVar v => TVars.add (v, b) | _ => I)) tvars tfrees); val cinstT = TVars.map (K (Thm.ctyp_of ctxt)) instT; val cinst = Vars.build (fold2 (fn v => fn t => (case v of Var (xi, T) => Vars.add ((xi, Term_Subst.instantiateT instT T), Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) | _ => I)) vars frees); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; -fun locale_target_declaration locale syntax decl lthy = lthy +fun locale_target_declaration locale params decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> - Locale.add_declaration locale syntax + Locale.add_declaration locale params (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = - locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) + locale_target_declaration locale {syntax = true, pos = Binding.pos_of b} + (const_decl phi_pred prmode ((b, mx), rhs)); (** locale operations **) -fun locale_declaration locale {syntax, pervasive} decl = +fun locale_declaration locale {syntax, pervasive, pos} decl = pervasive ? background_declaration decl - #> locale_target_declaration locale syntax decl + #> locale_target_declaration locale {syntax = syntax, pos = pos} decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/Isar/isar_cmd.ML b/src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML +++ b/src/Pure/Isar/isar_cmd.ML @@ -1,301 +1,302 @@ (* Title: Pure/Isar/isar_cmd.ML Author: Markus Wenzel, TU Muenchen Miscellaneous Isar commands. *) signature ISAR_CMD = sig val setup: Input.source -> theory -> theory val local_setup: Input.source -> Proof.context -> Proof.context val parse_ast_translation: Input.source -> theory -> theory val parse_translation: Input.source -> theory -> theory val print_translation: Input.source -> theory -> theory val typed_print_translation: Input.source -> theory -> theory val print_ast_translation: Input.source -> theory -> theory val translations: (xstring * string) Syntax.trrule list -> theory -> theory val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory val oracle: bstring * Position.range -> Input.source -> theory -> theory val declaration: {syntax: bool, pervasive: bool} -> Input.source -> local_theory -> local_theory val simproc_setup: string * Position.T -> string list -> Input.source -> local_theory -> local_theory val qed: Method.text_range option -> Toplevel.transition -> Toplevel.transition val terminal_proof: Method.text_range * Method.text_range option -> Toplevel.transition -> Toplevel.transition val default_proof: Toplevel.transition -> Toplevel.transition val immediate_proof: Toplevel.transition -> Toplevel.transition val done_proof: Toplevel.transition -> Toplevel.transition val skip_proof: Toplevel.transition -> Toplevel.transition val ml_diag: bool -> Input.source -> Toplevel.transition -> Toplevel.transition val diag_state: Proof.context -> Toplevel.state val diag_goal: Proof.context -> {context: Proof.context, facts: thm list, goal: thm} val pretty_theorems: bool -> Toplevel.state -> Pretty.T list val print_stmts: string list * (Facts.ref * Token.src list) list -> Toplevel.transition -> Toplevel.transition val print_thms: string list * (Facts.ref * Token.src list) list -> Toplevel.transition -> Toplevel.transition val print_prfs: bool -> string list * (Facts.ref * Token.src list) list option -> Toplevel.transition -> Toplevel.transition val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition val print_type: (string list * (string * string option)) -> Toplevel.transition -> Toplevel.transition end; structure Isar_Cmd: ISAR_CMD = struct (** theory declarations **) (* generic setup *) fun setup source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (" @ ML_Lex.read_source source @ ML_Lex.read ")") |> Context.theory_map; fun local_setup source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.local_setup (" @ ML_Lex.read_source source @ ML_Lex.read ")") |> Context.proof_map; (* translation functions *) fun parse_ast_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.parse_ast_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun parse_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.parse_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun print_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.print_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun typed_print_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.typed_print_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun print_ast_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.print_ast_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; (* translation rules *) fun read_trrules thy raw_rules = let val ctxt = Proof_Context.init_global thy; val read_root = #1 o dest_Type o Proof_Context.read_type_name {proper = true, strict = false} ctxt; in raw_rules |> map (Syntax.map_trrule (fn (r, s) => Syntax_Phases.parse_ast_pattern ctxt (read_root r, s))) end; fun translations args thy = Sign.add_trrules (read_trrules thy args) thy; fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy; (* oracles *) fun oracle (name, range) source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "val " @ ML_Lex.read_range range name @ ML_Lex.read (" = snd (Context.>>> (Context.map_theory_result (Thm.add_oracle (" ^ ML_Syntax.make_binding (name, #1 range) ^ ", ") @ ML_Lex.read_source source @ ML_Lex.read "))))") |> Context.theory_map; (* declarations *) fun declaration {syntax, pervasive} source = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Local_Theory.declaration {syntax = " ^ - Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ "} (") @ + Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ + ", pos = " ^ ML_Syntax.print_position (Position.thread_data ()) ^ "} (") @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.proof_map; (* simprocs *) fun simproc_setup name lhss source = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Simplifier.define_simproc_cmd (" ^ ML_Syntax.make_binding name ^ ") {lhss = " ^ ML_Syntax.print_strings lhss ^ ", proc = (") @ ML_Lex.read_source source @ ML_Lex.read ")})") |> Context.proof_map; (* local endings *) fun local_qed m = Toplevel.proof (Proof.local_qed (m, true)); val local_terminal_proof = Toplevel.proof o Proof.local_future_terminal_proof; val local_default_proof = Toplevel.proof Proof.local_default_proof; val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof; val local_done_proof = Toplevel.proof Proof.local_done_proof; val local_skip_proof = Toplevel.proof' Proof.local_skip_proof; (* global endings *) fun global_qed m = Toplevel.end_proof (K (Proof.global_qed (m, true))); val global_terminal_proof = Toplevel.end_proof o K o Proof.global_future_terminal_proof; val global_default_proof = Toplevel.end_proof (K Proof.global_default_proof); val global_immediate_proof = Toplevel.end_proof (K Proof.global_immediate_proof); val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof; val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof); (* common endings *) fun qed m = local_qed m o global_qed m; fun terminal_proof m = local_terminal_proof m o global_terminal_proof m; val default_proof = local_default_proof o global_default_proof; val immediate_proof = local_immediate_proof o global_immediate_proof; val done_proof = local_done_proof o global_done_proof; val skip_proof = local_skip_proof o global_skip_proof; (* diagnostic ML evaluation *) structure Diag_State = Proof_Data ( type T = Toplevel.state option; fun init _ = NONE; ); fun ml_diag verbose source = Toplevel.keep (fn state => let val opt_ctxt = try Toplevel.generic_theory_of state |> Option.map (Context.proof_of #> Diag_State.put (SOME state)); val flags = ML_Compiler.verbose verbose ML_Compiler.flags; in ML_Context.eval_source_in opt_ctxt flags source end); fun diag_state ctxt = (case Diag_State.get ctxt of SOME st => st | NONE => Toplevel.make_state NONE); val diag_goal = Proof.goal o Toplevel.proof_of o diag_state; val _ = Theory.setup (ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\state\) (Scan.succeed "Isar_Cmd.diag_state ML_context") #> ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\goal\) (Scan.succeed "Isar_Cmd.diag_goal ML_context")); (* theorems of theory or proof context *) fun pretty_theorems verbose st = if Toplevel.is_proof st then Proof_Context.pretty_local_facts verbose (Toplevel.context_of st) else let val ctxt = Toplevel.context_of st; val prev_thys = (case Toplevel.previous_theory_of st of SOME thy => [thy] | NONE => Theory.parents_of (Proof_Context.theory_of ctxt)); in Proof_Display.pretty_theorems_diff verbose prev_thys ctxt end; (* print theorems, terms, types etc. *) local fun string_of_stmts ctxt args = Attrib.eval_thms ctxt args |> map (Element.pretty_statement ctxt Thm.theoremK) |> Pretty.chunks2 |> Pretty.string_of; fun string_of_thms ctxt args = Pretty.string_of (Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args)); fun string_of_prfs full state arg = Pretty.string_of (case arg of NONE => let val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state); val thy = Proof_Context.theory_of ctxt; val prf = Thm.proof_of thm; val prop = Thm.full_prop_of thm; val prf' = Proofterm.rewrite_proof_notypes ([], []) prf; in Proof_Syntax.pretty_proof ctxt (if full then Proofterm.reconstruct_proof thy prop prf' else prf') end | SOME srcs => let val ctxt = Toplevel.context_of state; val pretty_proof = Proof_Syntax.pretty_standard_proof_of ctxt full; in Pretty.chunks (map pretty_proof (Attrib.eval_thms ctxt srcs)) end); fun string_of_prop ctxt s = let val prop = Syntax.read_prop ctxt s; val ctxt' = Proof_Context.augment prop ctxt; in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end; fun string_of_term ctxt s = let val t = Syntax.read_term ctxt s; val T = Term.type_of t; val ctxt' = Proof_Context.augment t ctxt; in Pretty.string_of (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)]) end; fun string_of_type ctxt (s, NONE) = let val T = Syntax.read_typ ctxt s in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end | string_of_type ctxt (s1, SOME s2) = let val ctxt' = Config.put show_sorts true ctxt; val raw_T = Syntax.parse_typ ctxt' s1; val S = Syntax.read_sort ctxt' s2; val T = Syntax.check_term ctxt' (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S))) |> Logic.dest_type; in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end; fun print_item string_of (modes, arg) = Toplevel.keep (fn state => Print_Mode.with_modes modes (fn () => writeln (string_of state arg)) ()); in val print_stmts = print_item (string_of_stmts o Toplevel.context_of); val print_thms = print_item (string_of_thms o Toplevel.context_of); val print_prfs = print_item o string_of_prfs; val print_prop = print_item (string_of_prop o Toplevel.context_of); val print_term = print_item (string_of_term o Toplevel.context_of); val print_type = print_item (string_of_type o Toplevel.context_of); end; end; diff --git a/src/Pure/Isar/local_theory.ML b/src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML +++ b/src/Pure/Isar/local_theory.ML @@ -1,440 +1,440 @@ (* Title: Pure/Isar/local_theory.ML Author: Makarius Local theory operations, with abstract target context. *) type local_theory = Proof.context; type generic_theory = Context.generic; structure Attrib = struct type binding = binding * Token.src list; type thms = (thm list * Token.src list) list; type fact = binding * thms; end; structure Locale = struct type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}; end; signature LOCAL_THEORY = sig type operations val assert: local_theory -> local_theory val level: Proof.context -> int val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory val background_naming_of: local_theory -> Name_Space.naming val map_background_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory val restore_background_naming: local_theory -> local_theory -> local_theory val full_name: local_theory -> binding -> string val new_group: local_theory -> local_theory val reset_group: local_theory -> local_theory val standard_morphism: local_theory -> Proof.context -> morphism val standard_morphism_theory: local_theory -> morphism val standard_form: local_theory -> Proof.context -> 'a Morphism.entity -> 'a val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val raw_theory: (theory -> theory) -> local_theory -> local_theory val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val background_theory: (theory -> theory) -> local_theory -> local_theory val target_of: local_theory -> Proof.context val target: (Proof.context -> Proof.context) -> local_theory -> local_theory val target_morphism: local_theory -> morphism val propagate_ml_env: generic_theory -> generic_theory val touch_ml_env: generic_theory -> generic_theory val operations_of: local_theory -> operations val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val notes_kind: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory - val declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration -> + val declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory val locale_dependency: Locale.registration -> local_theory -> local_theory val pretty: local_theory -> Pretty.T list val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory val set_defsort: sort -> local_theory -> local_theory val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> local_theory -> local_theory val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val type_alias: binding -> string -> local_theory -> local_theory val const_alias: binding -> string -> local_theory -> local_theory val init: {background_naming: Name_Space.naming, setup: theory -> Proof.context, conclude: local_theory -> Proof.context} -> operations -> theory -> local_theory val exit: local_theory -> Proof.context val exit_global: local_theory -> theory val exit_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * Proof.context val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory val begin_nested: local_theory -> Binding.scope * local_theory val end_nested: local_theory -> local_theory val end_nested_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * local_theory end; signature PRIVATE_LOCAL_THEORY = sig include LOCAL_THEORY val reset: local_theory -> local_theory end structure Local_Theory: PRIVATE_LOCAL_THEORY = struct (** local theory data **) (* type lthy *) type operations = {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory, notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory, abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory, - declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration_entity -> + declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration_entity -> local_theory -> local_theory, theory_registration: Locale.registration -> local_theory -> local_theory, locale_dependency: Locale.registration -> local_theory -> local_theory, pretty: local_theory -> Pretty.T list}; type lthy = {background_naming: Name_Space.naming, operations: operations, conclude: Proof.context -> Proof.context, target: Proof.context}; fun make_lthy (background_naming, operations, conclude, target) : lthy = {background_naming = background_naming, operations = operations, conclude = conclude, target = target}; (* context data *) structure Data = Proof_Data ( type T = lthy list; fun init _ = []; ); (* nested structure *) val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*) fun assert lthy = if level lthy = 0 then error "Missing local theory context" else lthy; fun assert_bottom lthy = let val _ = assert lthy; in if level lthy > 1 then error "Not at bottom of local theory nesting" else lthy end; fun assert_not_bottom lthy = let val _ = assert lthy; in if level lthy = 1 then error "Already at bottom of local theory nesting" else lthy end; val bottom_of = List.last o Data.get o assert; val top_of = hd o Data.get o assert; fun map_top f = assert #> Data.map (fn {background_naming, operations, conclude, target} :: parents => make_lthy (f (background_naming, operations, conclude, target)) :: parents); fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy); fun map_contexts f lthy = let val n = level lthy in lthy |> (Data.map o map_index) (fn (i, {background_naming, operations, conclude, target}) => make_lthy (background_naming, operations, conclude, target |> Context_Position.set_visible false |> f (n - i - 1) |> Context_Position.restore_visible target)) |> f n end; (* naming for background theory *) val background_naming_of = #background_naming o top_of; fun map_background_naming f = map_top (fn (background_naming, operations, conclude, target) => (f background_naming, operations, conclude, target)); val restore_background_naming = map_background_naming o K o background_naming_of; val full_name = Name_Space.full_name o background_naming_of; val new_group = map_background_naming Name_Space.new_group; val reset_group = map_background_naming Name_Space.reset_group; (* standard morphisms *) fun standard_morphism lthy ctxt = Morphism.set_context' lthy (Proof_Context.export_morphism lthy ctxt $> Morphism.thm_morphism' "Local_Theory.standard" (Goal.norm_result o Proof_Context.init_global) $> Morphism.binding_morphism "Local_Theory.standard_binding" (Name_Space.transform_binding (Proof_Context.naming_of lthy))); fun standard_morphism_theory lthy = standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); fun standard_form lthy ctxt x = Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); (* background theory *) fun raw_theory_result f lthy = let val (res, thy') = f (Proof_Context.theory_of lthy); val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; in (res, lthy') end; fun raw_theory f = #2 o raw_theory_result (f #> pair ()); fun background_theory_result f lthy = let val naming = background_naming_of lthy |> Name_Space.transform_naming (Proof_Context.naming_of lthy); in lthy |> raw_theory_result (fn thy => thy |> Sign.map_naming (K naming) |> f ||> Sign.restore_naming thy) end; fun background_theory f = #2 o background_theory_result (f #> pair ()); (* target contexts *) val target_of = #target o bottom_of; fun target f lthy = let val ctxt = target_of lthy; val ctxt' = ctxt |> Context_Position.set_visible false |> f |> Context_Position.restore_visible ctxt; val thy' = Proof_Context.theory_of ctxt'; in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; fun target_morphism lthy = standard_morphism lthy (target_of lthy); fun propagate_ml_env (context as Context.Proof lthy) = let val inherit = ML_Env.inherit [context] in lthy |> background_theory (Context.theory_map inherit) |> map_contexts (K (Context.proof_map inherit)) |> Context.Proof end | propagate_ml_env context = context; fun touch_ml_env context = if Context.enabled_tracing () then (case context of Context.Theory _ => ML_Env.touch context | Context.Proof _ => context) else context; (** operations **) val operations_of = #operations o top_of; fun operation f lthy = f (operations_of lthy) lthy; fun operation1 f x = operation (fn ops => f ops x); fun operation2 f x y = operation (fn ops => f ops x y); (* primitives *) val pretty = operation #pretty; val abbrev = operation2 #abbrev; val define = operation2 #define false; val define_internal = operation2 #define true; val notes_kind = operation2 #notes; fun declaration args = operation2 #declaration args o Morphism.entity; val theory_registration = operation1 #theory_registration; fun locale_dependency registration = assert_bottom #> operation1 #locale_dependency registration; (* theorems *) val notes = notes_kind ""; fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; fun add_thms_dynamic (binding, f) lthy = lthy |> background_theory_result (fn thy => thy |> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f)) |-> (fn name => map_contexts (fn _ => fn ctxt => Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #> - declaration {syntax = false, pervasive = false} (fn phi => + declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} (fn phi => let val binding' = Morphism.binding phi binding in Context.mapping (Global_Theory.alias_fact binding' name) (Proof_Context.alias_fact binding' name) end)); (* default sort *) fun set_defsort S = - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); (* syntax *) fun gen_syntax prep_type add mode raw_args lthy = let val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args; val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn _ => Proof_Context.generic_syntax add mode args') lthy end; val syntax = gen_syntax (K I); val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax; (* notation *) local fun gen_type_notation prep_type add mode raw_args lthy = let val prepare = prep_type lthy #> Logic.type_map (Assumption.export_term lthy (target_of lthy)); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (Proof_Context.generic_type_notation add mode args') lthy end; fun gen_notation prep_const add mode raw_args lthy = let val prepare = prep_const lthy #> Assumption.export_term lthy (target_of lthy); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (Proof_Context.generic_notation add mode args') lthy end; in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* name space aliases *) fun syntax_alias global_alias local_alias b name = - declaration {syntax = true, pervasive = false} (fn phi => + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn phi => let val b' = Morphism.binding phi b in Context.mapping (global_alias b' name) (local_alias b' name) end); val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias; val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias; (** manage targets **) (* main target *) fun init_target background_naming conclude operations target = Data.map (K [make_lthy (background_naming, operations, conclude, target)]) target fun init {background_naming, setup, conclude} operations thy = thy |> Sign.change_begin |> setup |> init_target background_naming (conclude #> target_of #> Sign.change_end_local) operations; val exit_of = #conclude o bottom_of; fun exit lthy = exit_of lthy (assert_bottom lthy); val exit_global = Proof_Context.theory_of o exit; fun exit_result decl (x, lthy) = let val ctxt = exit lthy; val phi = standard_morphism lthy ctxt; in (decl phi x, ctxt) end; fun exit_result_global decl (x, lthy) = let val thy = exit_global lthy; val thy_ctxt = Proof_Context.init_global thy; val phi = standard_morphism lthy thy_ctxt; in (decl phi x, thy) end; (* nested targets *) fun begin_nested lthy = let val _ = assert lthy; val (scope, target) = Proof_Context.new_scope lthy; val entry = make_lthy (background_naming_of lthy, operations_of lthy, Proof_Context.restore_naming lthy, target); val lthy' = Data.map (cons entry) target; in (scope, lthy') end; fun end_nested lthy = let val _ = assert_not_bottom lthy; val ({conclude, ...} :: rest) = Data.get lthy; in lthy |> Data.put rest |> reset |> conclude end; fun end_nested_result decl (x, lthy) = let val outer_lthy = end_nested lthy; val phi = Proof_Context.export_morphism lthy outer_lthy; in (decl phi x, outer_lthy) end; end; diff --git a/src/Pure/Isar/locale.ML b/src/Pure/Isar/locale.ML --- a/src/Pure/Isar/locale.ML +++ b/src/Pure/Isar/locale.ML @@ -1,803 +1,804 @@ (* Title: Pure/Isar/locale.ML Author: Clemens Ballarin, TU Muenchen Locales -- managed Isar proof contexts, based on Pure predicates. Draws basic ideas from Florian Kammueller's original version of locales, but uses the richer infrastructure of Isar instead of the raw meta-logic. Furthermore, structured composition of contexts (with merge and instantiation) is provided, as well as type-inference of the signature parts and predicate definitions of the specification text. Interpretation enables the transfer of declarations and theorems to other contexts, namely those defined by theories, structured proofs and locales themselves. A comprehensive account of locales is available: [1] Clemens Ballarin. Locales: a module system for mathematical theories. Journal of Automated Reasoning, 52(2):123-153, 2014. See also: [2] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar. In Stefano Berardi et al., Types for Proofs and Programs: International Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004. [3] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing Dependencies between Locales. Technical Report TUM-I0607, Technische Universitaet Muenchen, 2006. [4] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108, pages 31-43, 2006. *) signature LOCALE = sig (* Locale specification *) val register_locale: binding -> (string * sort) list * ((string * typ) * mixfix) list -> term option * term list -> thm option * thm option -> thm list -> Element.context_i list -> Morphism.declaration_entity list -> (string * Attrib.fact list) list -> (string * morphism) list -> theory -> theory val locale_space: theory -> Name_Space.T val intern: theory -> xstring -> string val check: theory -> xstring * Position.T -> string val extern: theory -> string -> xstring val markup_name: Proof.context -> string -> string val pretty_name: Proof.context -> string -> Pretty.T val defined: theory -> string -> bool val parameters_of: theory -> string -> (string * sort) list * ((string * typ) * mixfix) list val params_of: theory -> string -> ((string * typ) * mixfix) list val intros_of: theory -> string -> thm option * thm option val axioms_of: theory -> string -> thm list val instance_of: theory -> string -> morphism -> term list val specification_of: theory -> string -> term option * term list val hyp_spec_of: theory -> string -> Element.context_i list (* Storing results *) val add_facts: string -> string -> Attrib.fact list -> Proof.context -> Proof.context - val add_declaration: string -> bool -> Morphism.declaration_entity -> Proof.context -> Proof.context + val add_declaration: string -> {syntax: bool, pos: Position.T} -> + Morphism.declaration_entity -> Proof.context -> Proof.context (* Activation *) val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic val activate_declarations: string * morphism -> Proof.context -> Proof.context val init: string -> theory -> Proof.context (* Reasoning about locales *) val get_witnesses: Proof.context -> thm list val get_intros: Proof.context -> thm list val get_unfolds: Proof.context -> thm list val witness_add: attribute val intro_add: attribute val unfold_add: attribute val intro_locales_tac: {strict: bool, eager: bool} -> Proof.context -> thm list -> tactic (* Registrations and dependencies *) type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism} val amend_registration: registration -> Context.generic -> Context.generic val add_registration: registration -> Context.generic -> Context.generic val registrations_of: Context.generic -> string -> (string * morphism) list val add_dependency: string -> registration -> theory -> theory (* Diagnostic *) val get_locales: theory -> string list val locale_notes: theory -> string -> (string * Attrib.fact list) list val pretty_locales: theory -> bool -> Pretty.T val pretty_locale: theory -> bool -> string -> Pretty.T val pretty_registrations: Proof.context -> string -> Pretty.T val pretty_locale_deps: theory -> {name: string, parents: string list, body: Pretty.T} list type locale_dependency = {source: string, target: string, prefix: (string * bool) list, morphism: morphism, pos: Position.T, serial: serial} val dest_dependencies: theory list -> theory -> locale_dependency list val tracing : Proof.context -> string -> unit end; structure Locale: LOCALE = struct datatype ctxt = datatype Element.ctxt; (*** Locales ***) type dep = {name: string, morphisms: morphism * morphism, pos: Position.T, serial: serial}; fun eq_dep (dep1: dep, dep2: dep) = #serial dep1 = #serial dep2; fun transfer_dep thy ({name, morphisms, pos, serial}: dep) : dep = {name = name, morphisms = apply2 (Morphism.set_context thy) morphisms, pos = pos, serial = serial}; fun make_dep (name, morphisms) : dep = {name = name, morphisms = apply2 Morphism.reset_context morphisms, pos = Position.thread_data (), serial = serial ()}; (*table of mixin lists, per list mixins in reverse order of declaration; lists indexed by registration/dependency serial, entries for empty lists may be omitted*) type mixin = (morphism * bool) * serial; type mixins = mixin list Inttab.table; fun lookup_mixins (mixins: mixins) serial' = Inttab.lookup_list mixins serial'; val merge_mixins: mixins * mixins -> mixins = Inttab.merge_list (eq_snd op =); fun insert_mixin serial' (morph, b) : mixins -> mixins = Inttab.cons_list (serial', ((Morphism.reset_context morph, b), serial ())); fun rename_mixin (old, new) (mixins: mixins) = (case Inttab.lookup mixins old of NONE => mixins | SOME mixin => Inttab.delete old mixins |> Inttab.update_new (new, mixin)); fun compose_mixins (mixins: mixin list) = fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity; datatype locale = Loc of { (* static part *) (*type and term parameters*) parameters: (string * sort) list * ((string * typ) * mixfix) list, (*assumptions (as a single predicate expression) and defines*) spec: term option * term list, intros: thm option * thm option, axioms: thm list, (*diagnostic device: theorem part of hypothetical body as specified by the user*) hyp_spec: Element.context_i list, (* dynamic part *) (*syntax declarations*) syntax_decls: (Morphism.declaration_entity * serial) list, (*theorem declarations*) notes: ((string * Attrib.fact list) * serial) list, (*locale dependencies (sublocale relation) in reverse order*) dependencies: dep list, (*mixin part of dependencies*) mixins: mixins }; fun mk_locale ((parameters, spec, intros, axioms, hyp_spec), ((syntax_decls, notes), (dependencies, mixins))) = Loc {parameters = parameters, spec = spec, intros = intros, axioms = axioms, hyp_spec = hyp_spec, syntax_decls = syntax_decls, notes = notes, dependencies = dependencies, mixins = mixins}; fun map_locale f (Loc {parameters, spec, intros, axioms, hyp_spec, syntax_decls, notes, dependencies, mixins}) = mk_locale (f ((parameters, spec, intros, axioms, hyp_spec), ((syntax_decls, notes), (dependencies, mixins)))); fun merge_locale (Loc {parameters, spec, intros, axioms, hyp_spec, syntax_decls, notes, dependencies, mixins}, Loc {syntax_decls = syntax_decls', notes = notes', dependencies = dependencies', mixins = mixins', ...}) = mk_locale ((parameters, spec, intros, axioms, hyp_spec), ((merge (eq_snd op =) (syntax_decls, syntax_decls'), merge (eq_snd op =) (notes, notes')), (merge eq_dep (dependencies, dependencies'), (merge_mixins (mixins, mixins'))))); structure Locales = Theory_Data ( type T = locale Name_Space.table; val empty : T = Name_Space.empty_table Markup.localeN; val merge = Name_Space.join_tables (K merge_locale); ); val locale_space = Name_Space.space_of_table o Locales.get; val intern = Name_Space.intern o locale_space; fun check thy = #1 o Name_Space.check (Context.Theory thy) (Locales.get thy); val _ = Theory.setup (ML_Antiquotation.inline_embedded \<^binding>\locale\ (Args.theory -- Scan.lift Parse.embedded_position >> (ML_Syntax.print_string o uncurry check))); fun extern thy = Name_Space.extern (Proof_Context.init_global thy) (locale_space thy); fun markup_extern ctxt = Name_Space.markup_extern ctxt (locale_space (Proof_Context.theory_of ctxt)); fun markup_name ctxt name = markup_extern ctxt name |-> Markup.markup; fun pretty_name ctxt name = markup_extern ctxt name |> Pretty.mark_str; val get_locale = Name_Space.lookup o Locales.get; val defined = is_some oo get_locale; fun the_locale thy name = (case get_locale thy name of SOME (Loc loc) => loc | NONE => error ("Unknown locale " ^ quote name)); fun register_locale binding parameters spec intros axioms hyp_spec syntax_decls notes dependencies thy = thy |> Locales.map (Name_Space.define (Context.Theory thy) true (binding, mk_locale ((parameters, spec, (apply2 o Option.map) Thm.trim_context intros, map Thm.trim_context axioms, map Element.trim_context_ctxt hyp_spec), ((map (fn decl => (Morphism.entity_reset_context decl, serial ())) syntax_decls, map (fn (a, facts) => ((a, map Attrib.trim_context_fact facts), serial ())) notes), (map (fn (name, morph) => make_dep (name, (morph, Morphism.identity))) dependencies, Inttab.empty)))) #> snd); (* FIXME Morphism.identity *) fun change_locale name = Locales.map o Name_Space.map_table_entry name o map_locale o apsnd; (** Primitive operations **) fun parameters_of thy = #parameters o the_locale thy; val params_of = #2 oo parameters_of; fun intros_of thy = (apply2 o Option.map) (Thm.transfer thy) o #intros o the_locale thy; fun axioms_of thy = map (Thm.transfer thy) o #axioms o the_locale thy; fun instance_of thy name morph = params_of thy name |> map (Morphism.term (Morphism.set_context thy morph) o Free o #1); fun specification_of thy = #spec o the_locale thy; fun hyp_spec_of thy = map (Element.transfer_ctxt thy) o #hyp_spec o the_locale thy; fun dependencies_of thy = map (transfer_dep thy) o #dependencies o the_locale thy; fun mixins_of thy name serial = lookup_mixins (#mixins (the_locale thy name)) serial |> (map o apfst o apfst) (Morphism.set_context thy); (* Print instance and qualifiers *) fun pretty_reg_inst ctxt qs (name, ts) = let fun print_qual (qual, mandatory) = qual ^ (if mandatory then "" else "?"); fun prt_quals qs = Pretty.str (space_implode "." (map print_qual qs)); val prt_term = Pretty.quote o Syntax.pretty_term ctxt; fun prt_term' t = if Config.get ctxt show_types then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::", Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)] else prt_term t; fun prt_inst ts = Pretty.block (Pretty.breaks (pretty_name ctxt name :: map prt_term' ts)); in (case qs of [] => prt_inst ts | qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":", Pretty.brk 1, prt_inst ts]) end; fun pretty_reg ctxt export (name, morph) = let val thy = Proof_Context.theory_of ctxt; val morph' = morph $> export; val qs = Morphism.binding_prefix morph'; val ts = instance_of thy name morph'; in pretty_reg_inst ctxt qs (name, ts) end; (*** Identifiers: activated locales in theory or proof context ***) type idents = term list list Symtab.table; (* name ~> instance (grouped by name) *) val empty_idents : idents = Symtab.empty; val insert_idents = Symtab.insert_list (eq_list (op aconv)); val merge_idents = Symtab.merge_list (eq_list (op aconv)); fun redundant_ident thy idents (name, instance) = exists (fn pat => Pattern.matchess thy (pat, instance)) (Symtab.lookup_list idents name); structure Idents = Generic_Data ( type T = idents; val empty = empty_idents; val merge = merge_idents; ); (** Resolve locale dependencies in a depth-first fashion **) local val roundup_bound = 120; fun add thy depth stem export (name, morph) (deps, marked) = if depth > roundup_bound then error "Roundup bound exceeded (sublocale relation probably not terminating)." else let val instance = instance_of thy name (morph $> stem $> export); in if redundant_ident thy marked (name, instance) then (deps, marked) else let (*no inheritance of mixins, regardless of requests by clients*) val dependencies = dependencies_of thy name |> map (fn dep as {morphisms = (morph', export'), ...} => (#name dep, morph' $> export' $> compose_mixins (mixins_of thy name (#serial dep)))); val marked' = insert_idents (name, instance) marked; val (deps', marked'') = fold_rev (add thy (depth + 1) (morph $> stem) export) dependencies ([], marked'); in ((name, morph $> stem) :: deps' @ deps, marked'') end end; in (* Note that while identifiers always have the external (exported) view, activate_dep is presented with the internal view. *) fun roundup thy activate_dep export (name, morph) (marked, input) = let (* Find all dependencies including new ones (which are dependencies enriching existing registrations). *) val (dependencies, marked') = add thy 0 Morphism.identity export (name, morph) ([], empty_idents); (* Filter out fragments from marked; these won't be activated. *) val dependencies' = filter_out (fn (name, morph) => redundant_ident thy marked (name, instance_of thy name (morph $> export))) dependencies; in (merge_idents (marked, marked'), input |> fold_rev activate_dep dependencies') end; end; (*** Registrations: interpretations in theories or proof contexts ***) val total_ident_ord = prod_ord fast_string_ord (list_ord Term_Ord.fast_term_ord); structure Idtab = Table(type key = string * term list val ord = total_ident_ord); type reg = {morphisms: morphism * morphism, pos: Position.T, serial: serial}; val eq_reg: reg * reg -> bool = op = o apply2 #serial; (* FIXME consolidate with locale dependencies, consider one data slot only *) structure Global_Registrations = Theory_Data' ( (*registrations, indexed by locale name and instance; unique registration serial points to mixin list*) type T = reg Idtab.table * mixins; val empty: T = (Idtab.empty, Inttab.empty); fun merge args = let val ctxt0 = Syntax.init_pretty_global (#1 (hd args)); fun recursive_merge ((regs1, mixins1), (regs2, mixins2)) : T = (Idtab.merge eq_reg (regs1, regs2), merge_mixins (mixins1, mixins2)) handle Idtab.DUP id => (*distinct interpretations with same base: merge their mixins*) let val reg1 = Idtab.lookup regs1 id |> the; val reg2 = Idtab.lookup regs2 id |> the; val reg2' = {morphisms = #morphisms reg2, pos = Position.thread_data (), serial = #serial reg1}; val regs2' = Idtab.update (id, reg2') regs2; val mixins2' = rename_mixin (#serial reg2, #serial reg1) mixins2; val _ = warning ("Removed duplicate interpretation after retrieving its mixins" ^ Position.here_list [#pos reg1, #pos reg2] ^ ":\n " ^ Pretty.string_of (pretty_reg_inst ctxt0 [] id)); in recursive_merge ((regs1, mixins1), (regs2', mixins2')) end; in Library.foldl1 recursive_merge (map #2 args) end; ); structure Local_Registrations = Proof_Data ( type T = Global_Registrations.T; val init = Global_Registrations.get; ); val get_registrations = Context.cases Global_Registrations.get Local_Registrations.get; fun map_registrations f (Context.Theory thy) = Context.Theory (Global_Registrations.map f thy) | map_registrations f (Context.Proof ctxt) = Context.Proof (Local_Registrations.map f ctxt); (* Primitive operations *) fun add_reg thy export (name, morph) = let val reg = {morphisms = (Morphism.reset_context morph, Morphism.reset_context export), pos = Position.thread_data (), serial = serial ()}; val id = (name, instance_of thy name (morph $> export)); in (map_registrations o apfst) (Idtab.insert (K false) (id, reg)) end; fun add_mixin serial' mixin = (* registration to be amended identified by its serial id *) (map_registrations o apsnd) (insert_mixin serial' mixin); val get_regs = #1 o get_registrations; fun get_mixins context (name, morph) = let val thy = Context.theory_of context; val (regs, mixins) = get_registrations context; in (case Idtab.lookup regs (name, instance_of thy name morph) of NONE => [] | SOME {serial, ...} => lookup_mixins mixins serial) end; fun collect_mixins context (name, morph) = let val thy = Context.theory_of context; in roundup thy (fn dep => fn mixins => merge (eq_snd op =) (mixins, get_mixins context dep)) Morphism.identity (name, morph) (insert_idents (name, instance_of thy name morph) empty_idents, []) |> snd |> filter (snd o fst) (* only inheritable mixins *) |> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph))) |> compose_mixins end; (*** Activate context elements of locale ***) fun activate_err msg kind (name, morph) context = cat_error msg ("The above error(s) occurred while activating " ^ kind ^ " of locale instance\n" ^ (pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |> Pretty.string_of)); fun init_element elem context = context |> Context.mapping I (Thm.unchecked_hyps #> Context_Position.not_really) |> Element.init elem |> Context.mapping I (fn ctxt => let val ctxt0 = Context.proof_of context in ctxt |> Context_Position.restore_visible ctxt0 |> Thm.restore_hyps ctxt0 end); (* Potentially lazy notes *) fun make_notes kind = map (fn ((b, atts), facts) => if null atts andalso forall (null o #2) facts then Lazy_Notes (kind, (b, Lazy.value (maps #1 facts))) else Notes (kind, [((b, atts), facts)])); fun locale_notes thy loc = fold (cons o #1) (#notes (the_locale thy loc)) []; fun lazy_notes thy loc = locale_notes thy loc |> maps (fn (kind, notes) => make_notes kind notes); fun consolidate_notes elems = elems |> map_filter (fn Lazy_Notes (_, (_, ths)) => SOME ths | _ => NONE) |> Lazy.consolidate |> ignore; fun force_notes (Lazy_Notes (kind, (b, ths))) = Notes (kind, [((b, []), [(Lazy.force ths, [])])]) | force_notes elem = elem; (* Declarations, facts and entire locale content *) val trace_locales = Attrib.setup_config_bool (Binding.make ("trace_locales", \<^here>)) (K false); fun tracing context msg = if Config.get context trace_locales then Output.tracing msg else (); fun trace kind (name, morph) context = tracing (Context.proof_of context) ("Activating " ^ kind ^ " of " ^ (pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |> Pretty.string_of)); fun activate_syntax_decls (name, morph) context = let val _ = trace "syntax" (name, morph) context; val thy = Context.theory_of context; val {syntax_decls, ...} = the_locale thy name; val form_syntax_decl = Morphism.form o Morphism.transform morph o Morphism.entity_set_context thy; in fold_rev (form_syntax_decl o #1) syntax_decls context handle ERROR msg => activate_err msg "syntax" (name, morph) context end; fun activate_notes activ_elem context export' (name, morph) input = let val thy = Context.theory_of context; val mixin = (case export' of NONE => Morphism.identity | SOME export => collect_mixins context (name, morph $> export) $> export); val morph' = Morphism.set_context thy (morph $> mixin); val notes' = map (Element.transform_ctxt morph') (lazy_notes thy name); in (notes', input) |-> fold (fn elem => fn res => activ_elem (Element.transfer_ctxt thy elem) res) end handle ERROR msg => activate_err msg "facts" (name, morph) context; fun activate_notes_trace activ_elem context export' (name, morph) context' = let val _ = trace "facts" (name, morph) context'; in activate_notes activ_elem context export' (name, morph) context' end; fun activate_all name thy activ_elem (marked, input) = let val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name; val input' = input |> (not (null params) ? activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |> (* FIXME type parameters *) (case asm of SOME A => activ_elem (Assumes [(Binding.empty_atts, [(A, [])])]) | _ => I) |> (not (null defs) ? activ_elem (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs))); val activate = activate_notes activ_elem (Context.Theory thy) NONE; in roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input') end; (** Public activation functions **) fun activate_facts export dep context = context |> Context_Position.set_visible_generic false |> pair (Idents.get context) |> roundup (Context.theory_of context) (activate_notes_trace init_element context export) (Morphism.default export) dep |-> Idents.put |> Context_Position.restore_visible_generic context; fun activate_declarations dep = Context.proof_map (fn context => context |> Context_Position.set_visible_generic false |> pair (Idents.get context) |> roundup (Context.theory_of context) activate_syntax_decls Morphism.identity dep |-> Idents.put |> Context_Position.restore_visible_generic context); fun init name thy = let val context = Context.Proof (Proof_Context.init_global thy); val marked = Idents.get context; in context |> Context_Position.set_visible_generic false |> pair empty_idents |> activate_all name thy init_element |-> (fn marked' => Idents.put (merge_idents (marked, marked'))) |> Context_Position.restore_visible_generic context |> Context.proof_of end; (*** Add and extend registrations ***) type registration = Locale.registration; fun amend_registration {mixin = NONE, ...} context = context | amend_registration {inst = (name, morph), mixin = SOME mixin, export} context = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val regs = get_regs context; val base = instance_of thy name (morph $> export); val serial' = (case Idtab.lookup regs (name, base) of NONE => error ("No interpretation of locale " ^ quote (markup_name ctxt name) ^ " with\nparameter instantiation " ^ space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available") | SOME {serial = serial', ...} => serial'); in add_mixin serial' mixin context end; (* Note that a registration that would be subsumed by an existing one will not be generated, and it will not be possible to amend it. *) fun add_registration {inst = (name, base_morph), mixin, export} context = let val thy = Context.theory_of context; val pos_morph = Morphism.binding_morphism "position" (Binding.set_pos (Position.thread_data ())); val mix_morph = (case mixin of NONE => base_morph | SOME (mix, _) => base_morph $> mix); val inst = instance_of thy name mix_morph; val idents = Idents.get context; in if redundant_ident thy idents (name, inst) then context (* FIXME amend mixins? *) else (idents, context) (* add new registrations with inherited mixins *) |> roundup thy (add_reg thy export) export (name, mix_morph) |> #2 (* add mixin *) |> amend_registration {inst = (name, mix_morph), mixin = mixin, export = export} (* activate import hierarchy as far as not already active *) |> activate_facts (SOME export) (name, mix_morph $> pos_morph) end; (*** Dependencies ***) fun registrations_of context loc = Idtab.fold_rev (fn ((name, _), {morphisms, ...}) => name = loc ? cons (name, morphisms)) (get_regs context) [] (*with inherited mixins*) |> map (fn (name, (base, export)) => (name, base $> (collect_mixins context (name, base $> export)) $> export)); fun add_dependency loc {inst = (name, morph), mixin, export} thy = let val dep = make_dep (name, (morph, export)); val add_dep = apfst (cons dep) #> apsnd (case mixin of NONE => I | SOME mixin => insert_mixin (#serial dep) mixin); val thy' = change_locale loc (apsnd add_dep) thy; val context' = Context.Theory thy'; val (_, regs) = fold_rev (roundup thy' cons export) (registrations_of context' loc) (Idents.get context', []); in fold_rev (fn inst => Context.theory_map (add_registration {inst = inst, mixin = NONE, export = export})) regs thy' end; (*** Storing results ***) fun add_facts loc kind facts ctxt = if null facts then ctxt else let val stored_notes = ((kind, map Attrib.trim_context_fact facts), serial ()); val applied_notes = make_notes kind facts; fun apply_notes morph = applied_notes |> fold (fn elem => fn thy => let val elem' = Element.transform_ctxt (Morphism.set_context thy morph) elem in Context.theory_map (Element.init elem') thy end); fun apply_registrations thy = fold_rev (apply_notes o #2) (registrations_of (Context.Theory thy) loc) thy; in ctxt |> Attrib.local_notes kind facts |> #2 |> Proof_Context.background_theory ((change_locale loc o apfst o apsnd) (cons stored_notes) #> apply_registrations) end; -fun add_declaration loc syntax decl = +fun add_declaration loc {syntax, pos} decl = let val decl0 = Morphism.entity_reset_context decl in syntax ? Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl0, serial ()))) - #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration decl0)] + #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration pos decl0)] end; (*** Reasoning about locales ***) (* Storage for witnesses, intro and unfold rules *) structure Thms = Generic_Data ( type T = thm Item_Net.T * thm Item_Net.T * thm Item_Net.T; val empty = (Thm.item_net, Thm.item_net, Thm.item_net); fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) = (Item_Net.merge (witnesses1, witnesses2), Item_Net.merge (intros1, intros2), Item_Net.merge (unfolds1, unfolds2)); ); fun get_thms which ctxt = map (Thm.transfer' ctxt) (which (Thms.get (Context.Proof ctxt))); val get_witnesses = get_thms (Item_Net.content o #1); val get_intros = get_thms (Item_Net.content o #2); val get_unfolds = get_thms (Item_Net.content o #3); val witness_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Item_Net.update (Thm.trim_context th) x, y, z))); val intro_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Item_Net.update (Thm.trim_context th) y, z))); val unfold_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Item_Net.update (Thm.trim_context th) z))); (* Tactics *) fun intro_locales_tac {strict, eager} ctxt = (if strict then Method.intros_tac else Method.try_intros_tac) ctxt (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else [])); val _ = Theory.setup (Method.setup \<^binding>\intro_locales\ (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = false})) "back-chain introduction rules of locales without unfolding predicates" #> Method.setup \<^binding>\unfold_locales\ (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = true})) "back-chain all introduction rules of locales"); (*** diagnostic commands and interfaces ***) fun get_locales thy = map #1 (Name_Space.dest_table (Locales.get thy)); fun pretty_locales thy verbose = Pretty.block (Pretty.breaks (Pretty.str "locales:" :: map (Pretty.mark_str o #1) (Name_Space.markup_table verbose (Proof_Context.init_global thy) (Locales.get thy)))); fun pretty_locale thy show_facts name = let val locale_ctxt = init name thy; fun cons_elem (elem as Notes _) = show_facts ? cons elem | cons_elem (elem as Lazy_Notes _) = show_facts ? cons elem | cons_elem elem = cons elem; val elems = activate_all name thy cons_elem (empty_idents, []) |> snd |> rev |> tap consolidate_notes |> map force_notes; in Pretty.block (Pretty.keyword1 "locale" :: Pretty.brk 1 :: pretty_name locale_ctxt name :: maps (fn elem => [Pretty.fbrk, Pretty.chunks (Element.pretty_ctxt locale_ctxt elem)]) elems) end; fun pretty_registrations ctxt name = (case registrations_of (Context.Proof ctxt) name of [] => Pretty.str "no interpretations" | regs => Pretty.big_list "interpretations:" (map (pretty_reg ctxt Morphism.identity) (rev regs))); fun pretty_locale_deps thy = let fun make_node name = {name = name, parents = map #name (dependencies_of thy name), body = pretty_locale thy false name}; val names = sort_strings (Name_Space.fold_table (cons o #1) (Locales.get thy) []); in map make_node names end; type locale_dependency = {source: string, target: string, prefix: (string * bool) list, morphism: morphism, pos: Position.T, serial: serial}; fun dest_dependencies prev_thys thy = let fun remove_prev loc prev_thy deps = (case get_locale prev_thy loc of NONE => deps | SOME (Loc {dependencies = prev_deps, ...}) => if eq_list eq_dep (prev_deps, deps) then [] else subtract eq_dep prev_deps deps); fun result loc (dep: dep) = let val morphism = op $> (#morphisms dep) in {source = #name dep, target = loc, prefix = Morphism.binding_prefix morphism, morphism = morphism, pos = #pos dep, serial = #serial dep} end; fun add (loc, Loc {dependencies = deps, ...}) = fold (cons o result loc) (fold (remove_prev loc) prev_thys deps); in Name_Space.fold_table add (Locales.get thy) [] |> sort (int_ord o apply2 #serial) end; end; diff --git a/src/Pure/Isar/method.ML b/src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML +++ b/src/Pure/Isar/method.ML @@ -1,836 +1,837 @@ (* Title: Pure/Isar/method.ML Author: Markus Wenzel, TU Muenchen Isar proof methods. *) signature METHOD = sig type method = thm list -> context_tactic val CONTEXT_METHOD: (thm list -> context_tactic) -> method val METHOD: (thm list -> tactic) -> method val fail: method val succeed: method val insert_tac: Proof.context -> thm list -> int -> tactic val insert: thm list -> method val SIMPLE_METHOD: tactic -> method val SIMPLE_METHOD': (int -> tactic) -> method val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method val goal_cases_tac: string list -> context_tactic val cheating: bool -> method val intro: Proof.context -> thm list -> method val elim: Proof.context -> thm list -> method val unfold: thm list -> Proof.context -> method val fold: thm list -> Proof.context -> method val atomize: bool -> Proof.context -> method val this: Proof.context -> method val fact: thm list -> Proof.context -> method val assm_tac: Proof.context -> int -> tactic val all_assm_tac: Proof.context -> tactic val assumption: Proof.context -> method val rule_trace: bool Config.T val trace: Proof.context -> thm list -> unit val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val intros_tac: Proof.context -> thm list -> thm list -> tactic val try_intros_tac: Proof.context -> thm list -> thm list -> tactic val rule: Proof.context -> thm list -> method val erule: Proof.context -> int -> thm list -> method val drule: Proof.context -> int -> thm list -> method val frule: Proof.context -> int -> thm list -> method val method_space: Context.generic -> Name_Space.T val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic val clean_facts: thm list -> thm list val set_facts: thm list -> Proof.context -> Proof.context val get_facts: Proof.context -> thm list type combinator_info val no_combinator_info: combinator_info datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list val map_source: (Token.src -> Token.src) -> text -> text val primitive_text: (Proof.context -> thm -> thm) -> text val succeed_text: text val standard_text: text val this_text: text val done_text: text val sorry_text: bool -> text val finish_text: text option * bool -> text val print_methods: bool -> Proof.context -> unit val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val check_text: Proof.context -> text -> text val checked_text: text -> bool val method_syntax: (Proof.context -> method) context_parser -> Token.src -> Proof.context -> method val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory val local_setup: binding -> (Proof.context -> method) context_parser -> string -> local_theory -> local_theory val method_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory val method: Proof.context -> Token.src -> Proof.context -> method val method_closure: Proof.context -> Token.src -> Token.src val closure: bool Config.T val method_cmd: Proof.context -> Token.src -> Proof.context -> method val detect_closure_state: thm -> bool val STATIC: (unit -> unit) -> context_tactic val RUNTIME: context_tactic -> context_tactic val sleep: Time.time -> context_tactic val evaluate: text -> Proof.context -> method val evaluate_runtime: text -> Proof.context -> method type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T} val modifier: attribute -> Position.T -> modifier val old_section_parser: bool Config.T val sections: modifier parser list -> unit context_parser type text_range = text * Position.range val text: text_range option -> text option val position: text_range option -> Position.T val reports_of: text_range -> Position.report list val report: text_range -> unit val parser: int -> text_range parser val parse: text_range parser val parse_by: ((text_range * text_range option) * Position.report list) parser val read: Proof.context -> Token.src -> text val read_closure: Proof.context -> Token.src -> text * Token.src val read_closure_input: Proof.context -> Input.source -> text * Token.src val text_closure: text context_parser end; structure Method: METHOD = struct (** proof methods **) (* type method *) type method = thm list -> context_tactic; fun CONTEXT_METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac) #> Seq.maps_results (tac facts); fun METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac THEN tac facts); val fail = METHOD (K no_tac); val succeed = METHOD (K all_tac); (* insert facts *) fun insert_tac _ [] _ = all_tac | insert_tac ctxt facts i = EVERY (map (fn r => resolve_tac ctxt [Thm.forall_intr_vars r COMP_INCR revcut_rl] i) facts); fun insert thms = CONTEXT_METHOD (fn _ => fn (ctxt, st) => st |> ALLGOALS (insert_tac ctxt thms) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> (ALLGOALS (insert_tac ctxt facts) THEN tac) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD'' quant tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> quant (insert_tac ctxt facts THEN' tac) |> TACTIC_CONTEXT ctxt); val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; (* goals as cases *) fun goal_cases_tac case_names : context_tactic = fn (ctxt, st) => let val cases = (if null case_names then map string_of_int (1 upto Thm.nprems_of st) else case_names) |> map (rpair [] o rpair []) |> Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params st)); in CONTEXT_CASES cases all_tac (ctxt, st) end; (* cheating *) fun cheating int = CONTEXT_METHOD (fn _ => fn (ctxt, st) => if int orelse Config.get ctxt quick_and_dirty then TACTIC_CONTEXT ctxt (ALLGOALS (Skip_Proof.cheat_tac ctxt) st) else error "Cheating requires quick_and_dirty mode!"); (* unfold intro/elim rules *) fun intro ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ctxt ths)); fun elim ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt ths)); (* unfold/fold definitions *) fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); (* atomize rule statements *) fun atomize false ctxt = SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt) | atomize true ctxt = Context_Tactic.CONTEXT_TACTIC o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt)); (* this -- resolve facts directly *) fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single)); (* fact -- composition by facts from context *) fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt) | fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules); (* assumption *) local fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) => if cond (Logic.strip_assums_concl prop) then resolve_tac ctxt [rule] i else no_tac); in fun assm_tac ctxt = assume_tac ctxt APPEND' Goal.assume_rule_tac ctxt APPEND' cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND' cond_rtac ctxt (can Logic.dest_term) Drule.termI; fun all_assm_tac ctxt = let fun tac i st = if i > Thm.nprems_of st then all_tac st else ((assm_tac ctxt i THEN tac i) ORELSE tac (i + 1)) st; in tac 1 end; fun assumption ctxt = METHOD (HEADGOAL o (fn [] => assm_tac ctxt | [fact] => solve_tac ctxt [fact] | _ => K no_tac)); fun finish immed ctxt = METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac ctxt)); end; (* rule etc. -- single-step refinements *) val rule_trace = Attrib.setup_config_bool \<^binding>\rule_trace\ (fn _ => false); fun trace ctxt rules = if Config.get ctxt rule_trace andalso not (null rules) then Pretty.big_list "rules:" (map (Thm.pretty_thm_item ctxt) rules) |> Pretty.string_of |> tracing else (); local fun gen_rule_tac tac ctxt rules facts = (fn i => fn st => if null facts then tac ctxt rules i st else Seq.maps (fn rule => (tac ctxt o single) rule i st) (Drule.multi_resolves (SOME ctxt) facts rules)) THEN_ALL_NEW Goal.norm_hhf_tac ctxt; fun gen_arule_tac tac ctxt j rules facts = EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j (assume_tac ctxt)); fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) => let val rules = if not (null arg_rules) then arg_rules else flat (Context_Rules.find_rules ctxt false facts goal); in trace ctxt rules; tac ctxt rules facts i end); fun meth tac x y = METHOD (HEADGOAL o tac x y); fun meth' tac x y z = METHOD (HEADGOAL o tac x y z); in val rule_tac = gen_rule_tac resolve_tac; val rule = meth rule_tac; val some_rule_tac = gen_some_rule_tac rule_tac; val some_rule = meth some_rule_tac; val erule = meth' (gen_arule_tac eresolve_tac); val drule = meth' (gen_arule_tac dresolve_tac); val frule = meth' (gen_arule_tac forward_tac); end; (* intros_tac -- pervasive search spanned by intro rules *) fun gen_intros_tac goals ctxt intros facts = goals (insert_tac ctxt facts THEN' REPEAT_ALL_NEW (resolve_tac ctxt intros)) THEN Tactic.distinct_subgoals_tac; val intros_tac = gen_intros_tac ALLGOALS; val try_intros_tac = gen_intros_tac TRYALL; (** method syntax **) (* context data *) structure Data = Generic_Data ( type T = {methods: ((Token.src -> Proof.context -> method) * string) Name_Space.table, ml_tactic: (morphism -> thm list -> tactic) option, facts: thm list option}; val empty : T = {methods = Name_Space.empty_table Markup.methodN, ml_tactic = NONE, facts = NONE}; fun merge ({methods = methods1, ml_tactic = ml_tactic1, facts = facts1}, {methods = methods2, ml_tactic = ml_tactic2, facts = facts2}) : T = {methods = Name_Space.merge_tables (methods1, methods2), ml_tactic = merge_options (ml_tactic1, ml_tactic2), facts = merge_options (facts1, facts2)}; ); fun map_data f = Data.map (fn {methods, ml_tactic, facts} => let val (methods', ml_tactic', facts') = f (methods, ml_tactic, facts) in {methods = methods', ml_tactic = ml_tactic', facts = facts'} end); val get_methods = #methods o Data.get; val ops_methods = {get_data = get_methods, put_data = fn methods => map_data (fn (_, ml_tactic, facts) => (methods, ml_tactic, facts))}; val method_space = Name_Space.space_of_table o get_methods; (* ML tactic *) fun set_tactic ml_tactic = map_data (fn (methods, _, facts) => (methods, SOME ml_tactic, facts)); fun the_tactic context = #ml_tactic (Data.get context) |> \<^if_none>\raise Fail "Undefined ML tactic"\; val parse_tactic = Scan.state :|-- (fn context => Scan.lift (Args.embedded_declaration (fn source => let val tac = context |> ML_Context.expression (Input.pos_of source) (ML_Lex.read "Context.>> (Method.set_tactic (fn morphism: Morphism.morphism => fn facts: thm list => (" @ ML_Lex.read_source source @ ML_Lex.read ")))") |> the_tactic; in Morphism.entity (fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi))) end)) >> (fn decl => Morphism.form_entity (the_tactic (Morphism.form decl context)))); (* method facts *) val clean_facts = filter_out Thm.is_dummy; fun set_facts facts = (Context.proof_map o map_data) (fn (methods, ml_tactic, _) => (methods, ml_tactic, SOME (clean_facts facts))); val get_facts_generic = these o #facts o Data.get; val get_facts = get_facts_generic o Context.Proof; val _ = Theory.setup (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", \<^here>), get_facts_generic)); (* method text *) datatype combinator_info = Combinator_Info of {keywords: Position.T list}; fun combinator_info keywords = Combinator_Info {keywords = keywords}; val no_combinator_info = combinator_info []; datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int; datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list; fun map_source f (Source src) = Source (f src) | map_source _ (Basic meth) = Basic meth | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts); fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r); val succeed_text = Basic (K succeed); val standard_text = Source (Token.make_src ("standard", Position.none) []); val this_text = Basic this; val done_text = Basic (K (SIMPLE_METHOD all_tac)); fun sorry_text int = Basic (fn _ => cheating int); fun finish_text (NONE, immed) = Basic (finish immed) | finish_text (SOME txt, immed) = Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]); (* method definitions *) fun print_methods verbose ctxt = let val meths = get_methods (Context.Proof ctxt); fun prt_meth (name, (_, "")) = Pretty.mark_str name | prt_meth (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table verbose ctxt meths))] |> Pretty.writeln_chunks end; (* define *) fun define_global binding meth comment = Entity.define_global ops_methods binding (meth, comment); fun define binding meth comment = Entity.define ops_methods binding (meth, comment); (* check *) fun check_name ctxt = let val context = Context.Proof ctxt in #1 o Name_Space.check context (get_methods context) end; fun check_src ctxt = #1 o Token.check_src ctxt (get_methods o Context.Proof); fun check_text ctxt (Source src) = Source (check_src ctxt src) | check_text _ (Basic m) = Basic m | check_text ctxt (Combinator (x, y, body)) = Combinator (x, y, map (check_text ctxt) body); fun checked_text (Source src) = Token.checked_src src | checked_text (Basic _) = true | checked_text (Combinator (_, _, body)) = forall checked_text body; val _ = Theory.setup (ML_Antiquotation.inline_embedded \<^binding>\method\ (Args.context -- Scan.lift Parse.embedded_position >> (ML_Syntax.print_string o uncurry check_name))); (* method setup *) fun method_syntax scan src ctxt : method = let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end; fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd; fun method_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Method.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (")" ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* prepare methods *) fun method ctxt = let val table = get_methods (Context.Proof ctxt) in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; fun method_closure ctxt src = let val src' = map Token.init_assignable src; val ctxt' = Context_Position.not_really ctxt; val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm)); in map Token.closure src' end; val closure = Config.declare_bool ("Method.closure", \<^here>) (K true); fun method_cmd ctxt = check_src ctxt #> Config.get ctxt closure ? method_closure ctxt #> method ctxt; (* static vs. runtime state *) fun detect_closure_state st = (case try Logic.dest_term (Thm.concl_of (perhaps (try Goal.conclude) st)) of NONE => false | SOME t => Term.is_dummy_pattern t); fun STATIC test : context_tactic = fn (ctxt, st) => if detect_closure_state st then (test (); Seq.single (Seq.Result (ctxt, st))) else Seq.empty; fun RUNTIME (tac: context_tactic) (ctxt, st) = if detect_closure_state st then Seq.empty else tac (ctxt, st); fun sleep t = RUNTIME (fn ctxt_st => (OS.Process.sleep t; Seq.single (Seq.Result ctxt_st))); (* evaluate method text *) local val op THEN = Seq.THEN; fun BYPASS_CONTEXT (tac: tactic) = fn result => (case result of Seq.Error _ => Seq.single result | Seq.Result (ctxt, st) => tac st |> TACTIC_CONTEXT ctxt); val preparation = BYPASS_CONTEXT (ALLGOALS Goal.conjunction_tac); fun RESTRICT_GOAL i n method = BYPASS_CONTEXT (PRIMITIVE (Goal.restrict i n)) THEN method THEN BYPASS_CONTEXT (PRIMITIVE (Goal.unrestrict i)); fun SELECT_GOAL method i = RESTRICT_GOAL i 1 method; fun (method1 THEN_ALL_NEW method2) i (result : context_state Seq.result) = (case result of Seq.Error _ => Seq.single result | Seq.Result (_, st) => result |> method1 i |> Seq.maps (fn result' => (case result' of Seq.Error _ => Seq.single result' | Seq.Result (_, st') => result' |> Seq.INTERVAL method2 i (i + Thm.nprems_of st' - Thm.nprems_of st)))) fun COMBINATOR1 comb [meth] = comb meth | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument"; fun combinator Then = Seq.EVERY | combinator Then_All_New = (fn [] => Seq.single | methods => preparation THEN (foldl1 (op THEN_ALL_NEW) (map SELECT_GOAL methods) 1)) | combinator Orelse = Seq.FIRST | combinator Try = COMBINATOR1 Seq.TRY | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1 | combinator (Select_Goals n) = COMBINATOR1 (fn method => preparation THEN RESTRICT_GOAL 1 n method); in fun evaluate text ctxt0 facts = let val ctxt = set_facts facts ctxt0; fun eval0 m = Seq.single #> Seq.maps_results (m facts); fun eval (Basic m) = eval0 (m ctxt) | eval (Source src) = eval0 (method_cmd ctxt src ctxt) | eval (Combinator (_, c, txts)) = combinator c (map eval txts); in eval text o Seq.Result end; end; fun evaluate_runtime text ctxt = let val text' = text |> (map_source o map o Token.map_facts) (fn SOME name => (case Proof_Context.lookup_fact ctxt name of SOME {dynamic = true, thms} => K thms | _ => I) | NONE => I); val ctxt' = Config.put closure false ctxt; in fn facts => RUNTIME (fn st => evaluate text' ctxt' facts st) end; (** concrete syntax **) (* type modifier *) type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}; fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos}; (* sections *) val old_section_parser = Config.declare_bool ("Method.old_section_parser", \<^here>) (K false); local fun thms ss = Scan.repeats (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm); fun app {init, attribute, pos = _} ths context = fold_map (Thm.apply_attribute attribute) ths (Context.map_proof init context); fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- (fn (m, ths) => Scan.succeed (swap (app m ths context)))); in fun old_sections ss = Scan.repeat (section ss) >> K (); end; local fun sect (modifier : modifier parser) = Scan.depend (fn context => Scan.ahead Parse.not_eof -- Scan.trace modifier -- Scan.repeat (Scan.unless modifier Parse.thm) >> (fn ((tok0, ({init, attribute, pos}, modifier_toks)), xthms) => let val decl = (case Token.get_value tok0 of SOME (Token.Declaration decl) => decl | _ => let val ctxt = Context.proof_of context; val prep_att = Attrib.check_src ctxt #> map (Token.assign NONE); val thms = map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms; val facts = Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)] |> map (fn (_, bs) => - ((Binding.empty, [Attrib.internal (K attribute)]), Attrib.trim_context_thms bs)); + ((Binding.empty, [Attrib.internal pos (K attribute)]), + Attrib.trim_context_thms bs)); val decl = Morphism.entity (fn phi => fn context => let val psi = Morphism.set_context'' context phi in context |> Context.mapping I init |> Attrib.generic_notes "" (Attrib.transform_facts psi facts) |> snd end); val modifier_report = (#1 (Token.range_of modifier_toks), Position.entity_markup Markup.method_modifierN ("", pos)); val _ = Context_Position.reports ctxt (modifier_report :: Token.reports_of_value tok0); val _ = Token.assign (SOME (Token.Declaration decl)) tok0; in decl end); in (Morphism.form decl context, decl) end)); in fun sections ss = Args.context :|-- (fn ctxt => if Config.get ctxt old_section_parser then old_sections ss else Scan.repeat (sect (Scan.first ss)) >> K ()); end; (* extra rule methods *) fun xrule_meth meth = Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >> (fn (n, ths) => fn ctxt => meth ctxt n ths); (* text range *) type text_range = text * Position.range; fun text NONE = NONE | text (SOME (txt, _)) = SOME txt; fun position NONE = Position.none | position (SOME (_, (pos, _))) = pos; (* reports *) local fun keyword_positions (Source _) = [] | keyword_positions (Basic _) = [] | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) = keywords @ maps keyword_positions texts; in fun reports_of ((text, (pos, _)): text_range) = (pos, Markup.language_method) :: maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs "")) (keyword_positions text); fun report text_range = if Context_Position.reports_enabled0 () then Position.reports (reports_of text_range) else (); end; (* parser *) local fun is_symid_meth s = s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s; in fun parser pri = let val meth_name = Parse.token Parse.name; fun meth5 x = (meth_name >> (Source o single) || Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok => Source (Token.make_src ("cartouche", Position.none) [tok])) || Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x and meth4 x = (meth5 -- Parse.position (Parse.$$$ "?") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) || meth5 -- Parse.position (Parse.$$$ "+") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) || meth5 -- (Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]")) >> (fn (m, (((_, pos1), n), (_, pos2))) => Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) || meth5) x and meth3 x = (meth_name ::: Parse.args1 is_symid_meth >> Source || meth4) x and meth2 x = (Parse.enum1_positions "," meth3 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x and meth1 x = (Parse.enum1_positions ";" meth2 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then_All_New, ms))) x and meth0 x = (Parse.enum1_positions "|" meth1 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x; val meth = nth [meth0, meth1, meth2, meth3, meth4, meth5] pri handle General.Subscript => raise Fail ("Bad method parser priority " ^ string_of_int pri); in Scan.trace meth >> (fn (m, toks) => (m, Token.range_of toks)) end; val parse = parser 4; end; val parse_by = Parse.$$$ "by" |-- parse -- Scan.option parse >> (fn (m1, m2) => ((m1, m2), maps reports_of (m1 :: the_list m2))); (* read method text *) fun read ctxt src = (case Scan.read Token.stopper (Parse.!!! (parser 0 --| Scan.ahead Parse.eof)) src of SOME (text, range) => if checked_text text then text else (report (text, range); check_text ctxt text) | NONE => error ("Failed to parse method" ^ Position.here (#1 (Token.range_of src)))); fun read_closure ctxt src0 = let val src1 = map Token.init_assignable src0; val text = read ctxt src1 |> map_source (method_closure ctxt); val src2 = map Token.closure src1; in (text, src2) end; fun read_closure_input ctxt = let val keywords = Keyword.no_major_keywords (Thy_Header.get_keywords' ctxt) in Parse.read_embedded ctxt keywords (Scan.many Token.not_eof) #> read_closure ctxt end; val text_closure = Args.context -- Scan.lift (Parse.token Parse.embedded) >> (fn (ctxt, tok) => (case Token.get_value tok of SOME (Token.Source src) => read ctxt src | _ => let val (text, src) = read_closure_input ctxt (Token.input_of tok); val _ = Token.assign (SOME (Token.Source src)) tok; in text end)); (* theory setup *) val _ = Theory.setup (setup \<^binding>\fail\ (Scan.succeed (K fail)) "force failure" #> setup \<^binding>\succeed\ (Scan.succeed (K succeed)) "succeed" #> setup \<^binding>\sleep\ (Scan.lift Parse.real >> (fn s => fn _ => fn _ => sleep (seconds s))) "succeed after delay (in seconds)" #> setup \<^binding>\-\ (Scan.succeed (K (SIMPLE_METHOD all_tac))) "insert current facts, nothing else" #> setup \<^binding>\goal_cases\ (Scan.lift (Scan.repeat Args.name_token) >> (fn names => fn _ => CONTEXT_METHOD (fn _ => fn (ctxt, st) => (case drop (Thm.nprems_of st) names of [] => NONE | bad => if detect_closure_state st then NONE else SOME (fn () => ("Excessive case name(s): " ^ commas_quote (map Token.content_of bad) ^ Position.here (#1 (Token.range_of bad))))) |> (fn SOME msg => Seq.single (Seq.Error msg) | NONE => goal_cases_tac (map Token.content_of names) (ctxt, st))))) "bind cases for goals" #> setup \<^binding>\subproofs\ (text_closure >> (Context_Tactic.SUBPROOFS ooo evaluate_runtime)) "apply proof method to subproofs with closed derivation" #> setup \<^binding>\insert\ (Attrib.thms >> (K o insert)) "insert theorems, ignoring facts" #> setup \<^binding>\intro\ (Attrib.thms >> (fn ths => fn ctxt => intro ctxt ths)) "repeatedly apply introduction rules" #> setup \<^binding>\elim\ (Attrib.thms >> (fn ths => fn ctxt => elim ctxt ths)) "repeatedly apply elimination rules" #> setup \<^binding>\unfold\ (Attrib.thms >> unfold_meth) "unfold definitions" #> setup \<^binding>\fold\ (Attrib.thms >> fold_meth) "fold definitions" #> setup \<^binding>\atomize\ (Scan.lift (Args.mode "full") >> atomize) "present local premises as object-level statements" #> setup \<^binding>\rule\ (Attrib.thms >> (fn ths => fn ctxt => some_rule ctxt ths)) "apply some intro/elim rule" #> setup \<^binding>\erule\ (xrule_meth erule) "apply rule in elimination manner (improper)" #> setup \<^binding>\drule\ (xrule_meth drule) "apply rule in destruct manner (improper)" #> setup \<^binding>\frule\ (xrule_meth frule) "apply rule in forward manner (improper)" #> setup \<^binding>\this\ (Scan.succeed this) "apply current facts as rules" #> setup \<^binding>\fact\ (Attrib.thms >> fact) "composition by facts from context" #> setup \<^binding>\assumption\ (Scan.succeed assumption) "proof by assumption, preferring facts" #> setup \<^binding>\rename_tac\ (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> (fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs)))) "rename parameters of goal" #> setup \<^binding>\rotate_tac\ (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >> (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i)))) "rotate assumptions of goal" #> setup \<^binding>\tactic\ (parse_tactic >> (K o METHOD)) "ML tactic as proof method" #> setup \<^binding>\raw_tactic\ (parse_tactic >> (fn tac => fn _ => Context_Tactic.CONTEXT_TACTIC o tac)) "ML tactic as raw proof method" #> setup \<^binding>\use\ (Attrib.thms -- (Scan.lift (Parse.$$$ "in") |-- text_closure) >> (fn (thms, text) => fn ctxt => fn _ => evaluate_runtime text ctxt thms)) "indicate method facts and context for method expression"); (*final declarations of this structure!*) val unfold = unfold_meth; val fold = fold_meth; end; val CONTEXT_METHOD = Method.CONTEXT_METHOD; val METHOD = Method.METHOD; val SIMPLE_METHOD = Method.SIMPLE_METHOD; val SIMPLE_METHOD' = Method.SIMPLE_METHOD'; val SIMPLE_METHOD'' = Method.SIMPLE_METHOD''; diff --git a/src/Pure/Isar/spec_rules.ML b/src/Pure/Isar/spec_rules.ML --- a/src/Pure/Isar/spec_rules.ML +++ b/src/Pure/Isar/spec_rules.ML @@ -1,189 +1,189 @@ (* Title: Pure/Isar/spec_rules.ML Author: Makarius Rules that characterize specifications, with optional name and rough classification. NB: In the face of arbitrary morphisms, the original shape of specifications may get lost. *) signature SPEC_RULES = sig datatype recursion = Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion val recursion_ord: recursion ord val encode_recursion: recursion XML.Encode.T datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown val rough_classification_ord: rough_classification ord val equational_primrec: string list -> rough_classification val equational_recdef: rough_classification val equational_primcorec: string list -> rough_classification val equational_corec: rough_classification val equational: rough_classification val is_equational: rough_classification -> bool val is_inductive: rough_classification -> bool val is_co_inductive: rough_classification -> bool val is_relational: rough_classification -> bool val is_unknown: rough_classification -> bool val encode_rough_classification: rough_classification XML.Encode.T type spec_rule = {pos: Position.T, name: string, rough_classification: rough_classification, terms: term list, rules: thm list} val get: Proof.context -> spec_rule list val get_global: theory -> spec_rule list val dest_theory: theory -> spec_rule list val retrieve: Proof.context -> term -> spec_rule list val retrieve_global: theory -> term -> spec_rule list val add: binding -> rough_classification -> term list -> thm list -> local_theory -> local_theory val add_global: binding -> rough_classification -> term list -> thm list -> theory -> theory end; structure Spec_Rules: SPEC_RULES = struct (* recursion *) datatype recursion = Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion; val recursion_index = fn Primrec _ => 0 | Recdef => 1 | Primcorec _ => 2 | Corec => 3 | Unknown_Recursion => 4; fun recursion_ord (Primrec Ts1, Primrec Ts2) = list_ord fast_string_ord (Ts1, Ts2) | recursion_ord (Primcorec Ts1, Primcorec Ts2) = list_ord fast_string_ord (Ts1, Ts2) | recursion_ord rs = int_ord (apply2 recursion_index rs); val encode_recursion = let open XML.Encode in variant [fn Primrec a => ([], list string a), fn Recdef => ([], []), fn Primcorec a => ([], list string a), fn Corec => ([], []), fn Unknown_Recursion => ([], [])] end; (* rough classification *) datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown; fun rough_classification_ord (Equational r1, Equational r2) = recursion_ord (r1, r2) | rough_classification_ord cs = int_ord (apply2 (fn Equational _ => 0 | Inductive => 1 | Co_Inductive => 2 | Unknown => 3) cs); val equational_primrec = Equational o Primrec; val equational_recdef = Equational Recdef; val equational_primcorec = Equational o Primcorec; val equational_corec = Equational Corec; val equational = Equational Unknown_Recursion; val is_equational = fn Equational _ => true | _ => false; val is_inductive = fn Inductive => true | _ => false; val is_co_inductive = fn Co_Inductive => true | _ => false; val is_relational = is_inductive orf is_co_inductive; val is_unknown = fn Unknown => true | _ => false; val encode_rough_classification = let open XML.Encode in variant [fn Equational r => ([], encode_recursion r), fn Inductive => ([], []), fn Co_Inductive => ([], []), fn Unknown => ([], [])] end; (* rules *) type spec_rule = {pos: Position.T, name: string, rough_classification: rough_classification, terms: term list, rules: thm list}; fun eq_spec (specs: spec_rule * spec_rule) = (op =) (apply2 #name specs) andalso is_equal (rough_classification_ord (apply2 #rough_classification specs)) andalso eq_list (op aconv) (apply2 #terms specs) andalso eq_list Thm.eq_thm_prop (apply2 #rules specs); fun map_spec_rules f ({pos, name, rough_classification, terms, rules}: spec_rule) : spec_rule = {pos = pos, name = name, rough_classification = rough_classification, terms = terms, rules = map f rules}; structure Data = Generic_Data ( type T = spec_rule Item_Net.T; val empty : T = Item_Net.init eq_spec #terms; val merge = Item_Net.merge; ); (* get *) fun get_generic imports context = let val thy = Context.theory_of context; val transfer = Global_Theory.transfer_theories thy; fun imported spec = imports |> exists (fn thy => Item_Net.member (Data.get (Context.Theory thy)) spec); in Item_Net.content (Data.get context) |> filter_out imported |> (map o map_spec_rules) transfer end; val get = get_generic [] o Context.Proof; val get_global = get_generic [] o Context.Theory; fun dest_theory thy = rev (get_generic (Theory.parents_of thy) (Context.Theory thy)); (* retrieve *) fun retrieve_generic context = Item_Net.retrieve (Data.get context) #> (map o map_spec_rules) (Thm.transfer'' context); val retrieve = retrieve_generic o Context.Proof; val retrieve_global = retrieve_generic o Context.Theory; (* add *) fun add b rough_classification terms rules lthy = let val n = length terms; val thms0 = map Thm.trim_context (map (Drule.mk_term o Thm.cterm_of lthy) terms @ rules); in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of b} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi; val pos = Position.thread_data (); val name = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding psi b); val (terms', rules') = chop n (Morphism.fact psi thms0) |>> map (Thm.term_of o Drule.dest_term); in context |> (Data.map o Item_Net.update) {pos = pos, name = name, rough_classification = rough_classification, terms = terms', rules = rules'} end) end; fun add_global b rough_classification terms rules thy = thy |> (Context.theory_map o Data.map o Item_Net.update) {pos = Position.thread_data (), name = Sign.full_name thy b, rough_classification = rough_classification, terms = terms, rules = map Thm.trim_context rules}; end; diff --git a/src/Pure/ex/Def.thy b/src/Pure/ex/Def.thy --- a/src/Pure/ex/Def.thy +++ b/src/Pure/ex/Def.thy @@ -1,106 +1,106 @@ (* Title: Pure/ex/Def.thy Author: Makarius Primitive constant definition, without fact definition; automatic expansion via Simplifier (simproc). *) theory Def imports Pure keywords "def" :: thy_defn begin ML \ signature DEF = sig val get_def: Proof.context -> cterm -> thm option val def: (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term -> local_theory -> term * local_theory val def_cmd: (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string -> local_theory -> term * local_theory end; structure Def: DEF = struct (* context data *) type def = {lhs: term, eq: thm}; val eq_def : def * def -> bool = op aconv o apply2 #lhs; fun transform_def phi ({lhs, eq}: def) = {lhs = Morphism.term phi lhs, eq = Morphism.thm phi eq}; structure Data = Generic_Data ( type T = def Item_Net.T; val empty : T = Item_Net.init eq_def (single o #lhs); val merge = Item_Net.merge; ); fun declare_def lhs eq lthy = let val def0: def = {lhs = lhs, eq = Thm.trim_context eq} in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi in (Data.map o Item_Net.update) (transform_def psi def0) context end) end; fun get_def ctxt ct = let val thy = Proof_Context.theory_of ctxt; val data = Data.get (Context.Proof ctxt); val t = Thm.term_of ct; fun match_def {lhs, eq} = if Pattern.matches thy (lhs, t) then let val inst = Thm.match (Thm.cterm_of ctxt lhs, ct) in SOME (Thm.instantiate inst (Thm.transfer thy eq)) end else NONE; in Item_Net.retrieve_matching data t |> get_first match_def end; (* simproc setup *) val _ = (Theory.setup o Named_Target.theory_map) (Simplifier.define_simproc \<^binding>\expand_def\ {lhss = [Free ("x", TFree ("'a", []))], proc = K get_def}); (* Isar command *) fun gen_def prep_spec raw_var raw_params raw_spec lthy = let val ((vars, xs, get_pos, spec), _) = lthy |> prep_spec (the_list raw_var) raw_params [] raw_spec; val (((x, _), rhs), prove) = Local_Defs.derived_def lthy get_pos {conditional = false} spec; val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val ((lhs, (_, eq)), lthy') = lthy |> Local_Theory.define_internal ((b, mx), (Binding.empty_atts, rhs)); (*sanity check for original specification*) val _: thm = prove lthy' eq; in (lhs, declare_def lhs eq lthy') end; val def = gen_def Specification.check_spec_open; val def_cmd = gen_def Specification.read_spec_open; val _ = Outer_Syntax.local_theory \<^command_keyword>\def\ "primitive constant definition, without fact definition" (Scan.option Parse_Spec.constdecl -- Parse.prop -- Parse.for_fixes >> (fn ((decl, spec), params) => #2 o def_cmd decl params spec)); end; \ end diff --git a/src/Pure/simplifier.ML b/src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML +++ b/src/Pure/simplifier.ML @@ -1,433 +1,434 @@ (* Title: Pure/simplifier.ML Author: Tobias Nipkow and Markus Wenzel, TU Muenchen Generic simplifier, suitable for most logics (see also raw_simplifier.ML for the actual meta-level rewriting engine). *) signature BASIC_SIMPLIFIER = sig include BASIC_RAW_SIMPLIFIER val simp_tac: Proof.context -> int -> tactic val asm_simp_tac: Proof.context -> int -> tactic val full_simp_tac: Proof.context -> int -> tactic val asm_lr_simp_tac: Proof.context -> int -> tactic val asm_full_simp_tac: Proof.context -> int -> tactic val safe_simp_tac: Proof.context -> int -> tactic val safe_asm_simp_tac: Proof.context -> int -> tactic val safe_full_simp_tac: Proof.context -> int -> tactic val safe_asm_lr_simp_tac: Proof.context -> int -> tactic val safe_asm_full_simp_tac: Proof.context -> int -> tactic val simplify: Proof.context -> thm -> thm val asm_simplify: Proof.context -> thm -> thm val full_simplify: Proof.context -> thm -> thm val asm_lr_simplify: Proof.context -> thm -> thm val asm_full_simplify: Proof.context -> thm -> thm end; signature SIMPLIFIER = sig include BASIC_SIMPLIFIER val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val attrib: (thm -> Proof.context -> Proof.context) -> attribute val simp_add: attribute val simp_del: attribute val simp_flip: attribute val cong_add: attribute val cong_del: attribute val check_simproc: Proof.context -> xstring * Position.T -> string val the_simproc: Proof.context -> string -> simproc type 'a simproc_spec = {lhss: 'a list, proc: morphism -> Proof.context -> cterm -> thm option} val make_simproc: Proof.context -> string -> term simproc_spec -> simproc val define_simproc: binding -> term simproc_spec -> local_theory -> local_theory val define_simproc_cmd: binding -> string simproc_spec -> local_theory -> local_theory val pretty_simpset: bool -> Proof.context -> Pretty.T val default_mk_sym: Proof.context -> thm -> thm option val prems_of: Proof.context -> thm list val add_simp: thm -> Proof.context -> Proof.context val del_simp: thm -> Proof.context -> Proof.context val init_simpset: thm list -> Proof.context -> Proof.context val add_eqcong: thm -> Proof.context -> Proof.context val del_eqcong: thm -> Proof.context -> Proof.context val add_cong: thm -> Proof.context -> Proof.context val del_cong: thm -> Proof.context -> Proof.context val add_prems: thm list -> Proof.context -> Proof.context val mksimps: Proof.context -> thm -> thm list val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_term_ord: term ord -> Proof.context -> Proof.context val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context type trace_ops val set_trace_ops: trace_ops -> theory -> theory val rewrite: Proof.context -> conv val asm_rewrite: Proof.context -> conv val full_rewrite: Proof.context -> conv val asm_lr_rewrite: Proof.context -> conv val asm_full_rewrite: Proof.context -> conv val cong_modifiers: Method.modifier parser list val simp_modifiers': Method.modifier parser list val simp_modifiers: Method.modifier parser list val method_setup: Method.modifier parser list -> theory -> theory val unsafe_solver_tac: Proof.context -> int -> tactic val unsafe_solver: solver val safe_solver_tac: Proof.context -> int -> tactic val safe_solver: solver end; structure Simplifier: SIMPLIFIER = struct open Raw_Simplifier; (** declarations **) (* attributes *) fun attrib f = Thm.declaration_attribute (map_ss o f); val simp_add = attrib add_simp; val simp_del = attrib del_simp; val simp_flip = attrib flip_simp; val cong_add = attrib add_cong; val cong_del = attrib del_cong; (** named simprocs **) structure Simprocs = Generic_Data ( type T = simproc Name_Space.table; val empty : T = Name_Space.empty_table "simproc"; fun merge data : T = Name_Space.merge_tables data; ); (* get simprocs *) val get_simprocs = Simprocs.get o Context.Proof; fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1; val the_simproc = Name_Space.get o get_simprocs; val _ = Theory.setup (ML_Antiquotation.value_embedded \<^binding>\simproc\ (Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, name) => "Simplifier.the_simproc ML_context " ^ ML_Syntax.print_string (check_simproc ctxt name)))); (* define simprocs *) type 'a simproc_spec = {lhss: 'a list, proc: morphism -> Proof.context -> cterm -> thm option}; fun make_simproc ctxt name {lhss, proc} = let val ctxt' = fold Proof_Context.augment lhss ctxt; val lhss' = Variable.export_terms ctxt' ctxt lhss; in cert_simproc (Proof_Context.theory_of ctxt) name {lhss = lhss', proc = Morphism.entity proc} end; local fun def_simproc prep b {lhss, proc} lthy = let val simproc = make_simproc lthy (Local_Theory.full_name lthy b) {lhss = prep lthy lhss, proc = proc}; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val b' = Morphism.binding phi b; - val simproc' = transform_simproc phi simproc; - in - context - |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) - |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) - end) + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of b} + (fn phi => fn context => + let + val b' = Morphism.binding phi b; + val simproc' = transform_simproc phi simproc; + in + context + |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) + |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) + end) end; in val define_simproc = def_simproc Syntax.check_terms; val define_simproc_cmd = def_simproc Syntax.read_terms; end; (** congruence rule to protect foundational terms of local definitions **) local fun add_foundation_cong (binding, (const, target_params)) gthy = if null target_params then gthy else let val thy = Context.theory_of gthy; val cong = list_comb (const, target_params) |> Logic.varify_global |> Thm.global_cterm_of thy |> Thm.reflexive |> Thm.close_derivation \<^here>; val cong_binding = Binding.qualify_name true binding "cong"; in gthy |> Attrib.generic_notes Thm.theoremK [((cong_binding, []), [([cong], [])])] |> #2 end; val _ = Theory.setup (Generic_Target.add_foundation_interpretation add_foundation_cong); in end; (** pretty_simpset **) fun pretty_simpset verbose ctxt = let val pretty_term = Syntax.pretty_term ctxt; val pretty_thm = Thm.pretty_thm ctxt; val pretty_thm_item = Thm.pretty_thm_item ctxt; fun pretty_simproc (name, lhss) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.fbrk :: Pretty.fbreaks (map (Pretty.item o single o pretty_term) lhss)); fun pretty_cong_name (const, name) = pretty_term ((if const then Const else Free) (name, dummyT)); fun pretty_cong (name, thm) = Pretty.block [pretty_cong_name name, Pretty.str ":", Pretty.brk 1, pretty_thm thm]; val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss (simpset_of ctxt); val simprocs = Name_Space.markup_entries verbose ctxt (Name_Space.space_of_table (get_simprocs ctxt)) procs; in [Pretty.big_list "simplification rules:" (map (pretty_thm_item o #2) simps), Pretty.big_list "simplification procedures:" (map pretty_simproc simprocs), Pretty.big_list "congruences:" (map pretty_cong congs), Pretty.strs ("loopers:" :: map quote loopers), Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers), Pretty.strs ("safe solvers:" :: map quote safe_solvers)] |> Pretty.chunks end; (** simplification tactics and rules **) fun solve_all_tac solvers ctxt = let val subgoal_tac = Raw_Simplifier.subgoal_tac (Raw_Simplifier.set_solvers solvers ctxt); val solve_tac = subgoal_tac THEN_ALL_NEW (K no_tac); in DEPTH_SOLVE (solve_tac 1) end; (*NOTE: may instantiate unknowns that appear also in other subgoals*) fun generic_simp_tac safe mode ctxt = let val loop_tac = Raw_Simplifier.loop_tac ctxt; val (unsafe_solvers, solvers) = Raw_Simplifier.solvers ctxt; val solve_tac = FIRST' (map (Raw_Simplifier.solver ctxt) (rev (if safe then solvers else unsafe_solvers))); fun simp_loop_tac i = Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ctxt i THEN (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); in PREFER_GOAL (simp_loop_tac 1) end; local fun simp rew mode ctxt thm = let val (unsafe_solvers, _) = Raw_Simplifier.solvers ctxt; val tacf = solve_all_tac (rev unsafe_solvers); fun prover s th = Option.map #1 (Seq.pull (tacf s th)); in rew mode prover ctxt thm end; in val simp_thm = simp Raw_Simplifier.rewrite_thm; val simp_cterm = simp Raw_Simplifier.rewrite_cterm; end; (* tactics *) val simp_tac = generic_simp_tac false (false, false, false); val asm_simp_tac = generic_simp_tac false (false, true, false); val full_simp_tac = generic_simp_tac false (true, false, false); val asm_lr_simp_tac = generic_simp_tac false (true, true, false); val asm_full_simp_tac = generic_simp_tac false (true, true, true); (*not totally safe: may instantiate unknowns that appear also in other subgoals*) val safe_simp_tac = generic_simp_tac true (false, false, false); val safe_asm_simp_tac = generic_simp_tac true (false, true, false); val safe_full_simp_tac = generic_simp_tac true (true, false, false); val safe_asm_lr_simp_tac = generic_simp_tac true (true, true, false); val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); (* conversions *) val simplify = simp_thm (false, false, false); val asm_simplify = simp_thm (false, true, false); val full_simplify = simp_thm (true, false, false); val asm_lr_simplify = simp_thm (true, true, false); val asm_full_simplify = simp_thm (true, true, true); val rewrite = simp_cterm (false, false, false); val asm_rewrite = simp_cterm (false, true, false); val full_rewrite = simp_cterm (true, false, false); val asm_lr_rewrite = simp_cterm (true, true, false); val asm_full_rewrite = simp_cterm (true, true, true); (** concrete syntax of attributes **) (* add / del *) val simpN = "simp"; val flipN = "flip" val congN = "cong"; val onlyN = "only"; val no_asmN = "no_asm"; val no_asm_useN = "no_asm_use"; val no_asm_simpN = "no_asm_simp"; val asm_lrN = "asm_lr"; (* simprocs *) local val add_del = (Args.del -- Args.colon >> K (op delsimprocs) || Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) >> (fn f => fn simproc => Morphism.entity (fn phi => Thm.declaration_attribute (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc])))))); in val simproc_att = (Args.context -- Scan.lift add_del) :|-- (fn (ctxt, decl) => Scan.repeat1 (Scan.lift (Args.named_attribute (decl o the_simproc ctxt o check_simproc ctxt)))) >> (fn atts => Thm.declaration_attribute (fn th => fold (fn att => Thm.attribute_declaration (Morphism.form att) th) atts)); end; (* conversions *) local fun conv_mode x = ((Args.parens (Args.$$$ no_asmN) >> K simplify || Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || Scan.succeed asm_full_simplify) |> Scan.lift) x; in val simplified = conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute ths (fn context => f ((if null ths then I else Raw_Simplifier.clear_simpset) (Context.proof_of context) addsimps ths))); end; (* setup attributes *) val _ = Theory.setup (Attrib.setup \<^binding>\simp\ (Attrib.add_del simp_add simp_del) "declaration of Simplifier rewrite rule" #> Attrib.setup \<^binding>\cong\ (Attrib.add_del cong_add cong_del) "declaration of Simplifier congruence rule" #> Attrib.setup \<^binding>\simproc\ simproc_att "declaration of simplification procedures" #> Attrib.setup \<^binding>\simplified\ simplified "simplified rule"); (** method syntax **) val cong_modifiers = [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>)]; val simp_modifiers = [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), Args.$$$ simpN -- Args.$$$ flipN -- Args.colon >> K (Method.modifier simp_flip \<^here>), Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}] @ cong_modifiers; val simp_modifiers' = [Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), Args.$$$ flipN -- Args.colon >> K (Method.modifier simp_flip \<^here>), Args.$$$ onlyN -- Args.colon >> K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}] @ cong_modifiers; val simp_options = (Args.parens (Args.$$$ no_asmN) >> K simp_tac || Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || Scan.succeed asm_full_simp_tac); fun simp_method more_mods meth = Scan.lift simp_options --| Method.sections (more_mods @ simp_modifiers') >> (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts)); (** setup **) fun method_setup more_mods = Method.setup \<^binding>\simp\ (simp_method more_mods (fn ctxt => fn tac => fn facts => HEADGOAL (Method.insert_tac ctxt facts THEN' (CHANGED_PROP oo tac) ctxt))) "simplification" #> Method.setup \<^binding>\simp_all\ (simp_method more_mods (fn ctxt => fn tac => fn facts => ALLGOALS (Method.insert_tac ctxt facts) THEN (CHANGED_PROP o PARALLEL_ALLGOALS o tac) ctxt)) "simplification (all goals)"; fun unsafe_solver_tac ctxt = FIRST' [resolve_tac ctxt (Drule.reflexive_thm :: Raw_Simplifier.prems_of ctxt), assume_tac ctxt]; val unsafe_solver = mk_solver "Pure unsafe" unsafe_solver_tac; (*no premature instantiation of variables during simplification*) fun safe_solver_tac ctxt = FIRST' [match_tac ctxt (Drule.reflexive_thm :: Raw_Simplifier.prems_of ctxt), eq_assume_tac]; val safe_solver = mk_solver "Pure safe" safe_solver_tac; val _ = Theory.setup (method_setup [] #> Context.theory_map (map_ss (fn ctxt => empty_simpset ctxt setSSolver safe_solver setSolver unsafe_solver |> set_subgoaler asm_simp_tac))); end; structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier; open Basic_Simplifier;