# HG changeset patch # User huffman # Date 1235690838 28800 # Node ID 258f9adfdda543dd81f79974b7b2b582f4087847 # Parent 243a05a67c4167eb1ddafb0f42f92c924c15f539# Parent cd3f37ba3e2530a1ee5ad39a29acd84d4b67ce8f merged diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOL/Library/Bit.thy --- a/src/HOL/Library/Bit.thy Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOL/Library/Bit.thy Thu Feb 26 15:27:18 2009 -0800 @@ -79,14 +79,8 @@ end -lemma bit_1_plus_1 [simp]: "1 + 1 = (0 :: bit)" - unfolding plus_bit_def by simp - -lemma bit_add_self [simp]: "x + x = (0 :: bit)" - by (cases x) simp_all - -lemma bit_add_self_left [simp]: "x + (x + y) = (y :: bit)" - by simp +lemma bit_add_self: "x + x = (0 :: bit)" + unfolding plus_bit_def by (simp split: bit.split) lemma bit_mult_eq_1_iff [simp]: "x * y = (1 :: bit) \ x = 1 \ y = 1" unfolding times_bit_def by (simp split: bit.split) diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOL/List.thy --- a/src/HOL/List.thy Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOL/List.thy Thu Feb 26 15:27:18 2009 -0800 @@ -1438,10 +1438,10 @@ apply (auto split:nat.split) done -lemma last_conv_nth: "xs\[] \ last xs = xs!(length xs - Suc 0)" +lemma last_conv_nth: "xs\[] \ last xs = xs!(length xs - 1)" by(induct xs)(auto simp:neq_Nil_conv) -lemma butlast_conv_take: "butlast xs = take (length xs - Suc 0) xs" +lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" by (induct xs, simp, case_tac xs, simp_all) @@ -1461,6 +1461,12 @@ declare take_Cons [simp del] and drop_Cons [simp del] +lemma take_1_Cons [simp]: "take 1 (x # xs) = [x]" + unfolding One_nat_def by simp + +lemma drop_1_Cons [simp]: "drop 1 (x # xs) = xs" + unfolding One_nat_def by simp + lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)" by(clarsimp simp add:neq_Nil_conv) @@ -1588,17 +1594,17 @@ done lemma butlast_take: - "n <= length xs ==> butlast (take n xs) = take (n - Suc 0) xs" + "n <= length xs ==> butlast (take n xs) = take (n - 1) xs" by (simp add: butlast_conv_take min_max.inf_absorb1 min_max.inf_absorb2) lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)" -by (simp add: butlast_conv_take drop_take) +by (simp add: butlast_conv_take drop_take add_ac) lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs" by (simp add: butlast_conv_take min_max.inf_absorb1) lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" -by (simp add: butlast_conv_take drop_take) +by (simp add: butlast_conv_take drop_take add_ac) lemma hd_drop_conv_nth: "\ xs \ []; n < length xs \ \ hd(drop n xs) = xs!n" by(simp add: hd_conv_nth) @@ -2458,7 +2464,7 @@ done lemma length_remove1: - "length(remove1 x xs) = (if x : set xs then length xs - Suc 0 else length xs)" + "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)" apply (induct xs) apply (auto dest!:length_pos_if_in_set) done diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOL/Nat.thy Thu Feb 26 15:27:18 2009 -0800 @@ -701,6 +701,9 @@ lemma Suc_pred [simp]: "n>0 ==> Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) +lemma Suc_diff_1 [simp]: "0 < n ==> Suc (n - 1) = n" +unfolding One_nat_def by (rule Suc_pred) + lemma nat_add_left_cancel_le [simp]: "(k + m \ k + n) = (m\(n::nat))" by (induct k) simp_all @@ -1135,7 +1138,7 @@ by (cases m) (auto intro: le_add1) text {* Lemma for @{text gcd} *} -lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = Suc 0 | m = 0" +lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = 1 | m = 0" apply (drule sym) apply (rule disjCI) apply (rule nat_less_cases, erule_tac [2] _) diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOLCF/Fixrec.thy --- a/src/HOLCF/Fixrec.thy Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOLCF/Fixrec.thy Thu Feb 26 15:27:18 2009 -0800 @@ -583,6 +583,20 @@ use "Tools/fixrec_package.ML" +setup {* FixrecPackage.setup *} + +setup {* + FixrecPackage.add_matchers + [ (@{const_name up}, @{const_name match_up}), + (@{const_name sinl}, @{const_name match_sinl}), + (@{const_name sinr}, @{const_name match_sinr}), + (@{const_name spair}, @{const_name match_spair}), + (@{const_name cpair}, @{const_name match_cpair}), + (@{const_name ONE}, @{const_name match_ONE}), + (@{const_name TT}, @{const_name match_TT}), + (@{const_name FF}, @{const_name match_FF}) ] +*} + hide (open) const return bind fail run cases end diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOLCF/Tools/domain/domain_axioms.ML --- a/src/HOLCF/Tools/domain/domain_axioms.ML Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOLCF/Tools/domain/domain_axioms.ML Thu Feb 26 15:27:18 2009 -0800 @@ -39,7 +39,7 @@ fun one_con (con,args) = foldr /\# (list_ccomb (%%:con, mapn (idxs (length args)) 1 args)) args; in ("copy_def", %%:(dname^"_copy") == - /\"f" (list_ccomb (%%:(dname^"_when"), map one_con cons))) end; + /\ "f" (list_ccomb (%%:(dname^"_when"), map one_con cons))) end; (* -- definitions concerning the constructors, discriminators and selectors - *) @@ -107,7 +107,7 @@ [when_def, copy_def] @ con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @ [take_def, finite_def]) -end; (* let *) +end; (* let (calc_axioms) *) fun infer_props thy = map (apsnd (FixrecPackage.legacy_infer_prop thy)); @@ -117,6 +117,14 @@ fun add_defs_i x = snd o (PureThy.add_defs false) (map (Thm.no_attributes o apfst Binding.name) x); fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy; +fun add_matchers (((dname,_),cons) : eq) thy = + let + val con_names = map fst cons; + val mat_names = map mat_name con_names; + fun qualify n = Sign.full_name thy (Binding.name n); + val ms = map qualify con_names ~~ map qualify mat_names; + in FixrecPackage.add_matchers ms thy end; + in (* local *) fun add_axioms (comp_dnam, eqs : eq list) thy' = let @@ -125,7 +133,7 @@ val x_name = idx_name dnames "x"; fun copy_app dname = %%:(dname^"_copy")`Bound 0; val copy_def = ("copy_def" , %%:(comp_dname^"_copy") == - /\"f"(mk_ctuple (map copy_app dnames))); + /\ "f"(mk_ctuple (map copy_app dnames))); val bisim_def = ("bisim_def",%%:(comp_dname^"_bisim")==mk_lam("R", let fun one_con (con,args) = let @@ -164,7 +172,8 @@ in thy |> Sign.add_path comp_dnam |> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else [])) |> Sign.parent_path -end; + |> fold add_matchers eqs +end; (* let (add_axioms) *) end; (* local *) end; (* struct *) diff -r cd3f37ba3e25 -r 258f9adfdda5 src/HOLCF/Tools/fixrec_package.ML --- a/src/HOLCF/Tools/fixrec_package.ML Thu Feb 26 22:13:01 2009 +0100 +++ b/src/HOLCF/Tools/fixrec_package.ML Thu Feb 26 15:27:18 2009 -0800 @@ -8,17 +8,20 @@ sig val legacy_infer_term: theory -> term -> term val legacy_infer_prop: theory -> term -> term + val add_fixrec: bool -> (Attrib.binding * string) list list -> theory -> theory val add_fixrec_i: bool -> ((binding * attribute list) * term) list list -> theory -> theory val add_fixpat: Attrib.binding * string list -> theory -> theory val add_fixpat_i: (binding * attribute list) * term list -> theory -> theory + val add_matchers: (string * string) list -> theory -> theory + val setup: theory -> theory end; structure FixrecPackage: FIXREC_PACKAGE = struct (* legacy type inference *) - +(* used by the domain package *) fun legacy_infer_term thy t = singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t); @@ -33,15 +36,41 @@ fun fixrec_eq_err thy s eq = fixrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); +(*************************************************************************) +(***************************** building types ****************************) +(*************************************************************************) + (* ->> is taken from holcf_logic.ML *) -(* TODO: fix dependencies so we can import HOLCFLogic here *) -infixr 6 ->>; -fun S ->> T = Type (@{type_name "->"},[S,T]); +fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]); + +infixr 6 ->>; val (op ->>) = cfunT; + +fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U) + | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []); + +fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U + | binder_cfun _ = []; + +fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U + | body_cfun T = T; -(* extern_name is taken from domain/library.ML *) -fun extern_name con = case Symbol.explode con of - ("o"::"p"::" "::rest) => implode rest - | _ => con; +fun strip_cfun T : typ list * typ = + (binder_cfun T, body_cfun T); + +fun maybeT T = Type(@{type_name "maybe"}, [T]); + +fun dest_maybeT (Type(@{type_name "maybe"}, [T])) = T + | dest_maybeT T = raise TYPE ("dest_maybeT", [T], []); + +fun tupleT [] = @{typ "unit"} + | tupleT [T] = T + | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts); + +fun matchT T = body_cfun T ->> maybeT (tupleT (binder_cfun T)); + +(*************************************************************************) +(***************************** building terms ****************************) +(*************************************************************************) val mk_trp = HOLogic.mk_Trueprop; @@ -52,30 +81,86 @@ fun chead_of (Const(@{const_name Rep_CFun},_)$f$t) = chead_of f | chead_of u = u; -(* these are helpful functions copied from HOLCF/domain/library.ML *) -fun %: s = Free(s,dummyT); -fun %%: s = Const(s,dummyT); -infix 0 ==; fun S == T = %%:"==" $ S $ T; -infix 1 ===; fun S === T = %%:"op =" $ S $ T; -infix 9 ` ; fun f ` x = %%:@{const_name Rep_CFun} $ f $ x; +fun capply_const (S, T) = + Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T)); + +fun cabs_const (S, T) = + Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T)); + +fun mk_capply (t, u) = + let val (S, T) = + case Term.fastype_of t of + Type(@{type_name "->"}, [S, T]) => (S, T) + | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]); + in capply_const (S, T) $ t $ u end; + +infix 0 ==; val (op ==) = Logic.mk_equals; +infix 1 ===; val (op ===) = HOLogic.mk_eq; +infix 9 ` ; val (op `) = mk_capply; + + +fun mk_cpair (t, u) = + let val T = Term.fastype_of t + val U = Term.fastype_of u + val cpairT = T ->> U ->> HOLogic.mk_prodT (T, U) + in Const(@{const_name cpair}, cpairT) ` t ` u end; + +fun mk_cfst t = + let val T = Term.fastype_of t; + val (U, _) = HOLogic.dest_prodT T; + in Const(@{const_name cfst}, T ->> U) ` t end; + +fun mk_csnd t = + let val T = Term.fastype_of t; + val (_, U) = HOLogic.dest_prodT T; + in Const(@{const_name csnd}, T ->> U) ` t end; + +fun mk_csplit t = + let val (S, TU) = dest_cfunT (Term.fastype_of t); + val (T, U) = dest_cfunT TU; + val csplitT = (S ->> T ->> U) ->> HOLogic.mk_prodT (S, T) ->> U; + in Const(@{const_name csplit}, csplitT) ` t end; (* builds the expression (LAM v. rhs) *) -fun big_lambda v rhs = %%:@{const_name Abs_CFun}$(Term.lambda v rhs); +fun big_lambda v rhs = + cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs; (* builds the expression (LAM v1 v2 .. vn. rhs) *) fun big_lambdas [] rhs = rhs | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs); (* builds the expression (LAM . rhs) *) -fun lambda_ctuple [] rhs = big_lambda (%:"unit") rhs +fun lambda_ctuple [] rhs = big_lambda (Free("unit", HOLogic.unitT)) rhs | lambda_ctuple (v::[]) rhs = big_lambda v rhs | lambda_ctuple (v::vs) rhs = - %%:@{const_name csplit}`(big_lambda v (lambda_ctuple vs rhs)); + mk_csplit (big_lambda v (lambda_ctuple vs rhs)); (* builds the expression *) -fun mk_ctuple [] = %%:"UU" +fun mk_ctuple [] = @{term "UU::unit"} | mk_ctuple (t::[]) = t -| mk_ctuple (t::ts) = %%:@{const_name cpair}`t`(mk_ctuple ts); +| mk_ctuple (t::ts) = mk_cpair (t, mk_ctuple ts); + +fun mk_return t = + let val T = Term.fastype_of t + in Const(@{const_name Fixrec.return}, T ->> maybeT T) ` t end; + +fun mk_bind (t, u) = + let val (T, mU) = dest_cfunT (Term.fastype_of u); + val bindT = maybeT T ->> (T ->> mU) ->> mU; + in Const(@{const_name Fixrec.bind}, bindT) ` t ` u end; + +fun mk_mplus (t, u) = + let val mT = Term.fastype_of t + in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end; + +fun mk_run t = + let val mT = Term.fastype_of t + val T = dest_maybeT mT + in Const(@{const_name Fixrec.run}, mT ->> T) ` t end; + +fun mk_fix t = + let val (T, _) = dest_cfunT (Term.fastype_of t) + in Const(@{const_name fix}, (T ->> T) ->> T) ` t end; (*************************************************************************) (************* fixed-point definitions and unfolding theorems ************) @@ -84,22 +169,21 @@ fun add_fixdefs eqs thy = let val (lhss,rhss) = ListPair.unzip (map dest_eqs eqs); - val fixpoint = %%:@{const_name fix}`lambda_ctuple lhss (mk_ctuple rhss); + val fixpoint = mk_fix (lambda_ctuple lhss (mk_ctuple rhss)); fun one_def (l as Const(n,T)) r = let val b = Sign.base_name n in (b, (b^"_def", l == 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 (%%:@{const_name cfst}`r) :: defs ls (%%:@{const_name csnd}`r); - val (names, pre_fixdefs) = ListPair.unzip (defs lhss fixpoint); + | defs (l::ls) r = one_def l (mk_cfst r) :: defs ls (mk_csnd r); + val (names, fixdefs) = ListPair.unzip (defs lhss fixpoint); - val fixdefs = map (apsnd (legacy_infer_prop thy)) pre_fixdefs; val (fixdef_thms, thy') = PureThy.add_defs false (map (Thm.no_attributes o apfst Binding.name) fixdefs) thy; val ctuple_fixdef_thm = foldr1 (fn (x,y) => @{thm cpair_equalI} OF [x,y]) fixdef_thms; - val ctuple_unfold = legacy_infer_term thy' (mk_trp (mk_ctuple lhss === mk_ctuple rhss)); + val ctuple_unfold = mk_trp (mk_ctuple lhss === mk_ctuple rhss); val ctuple_unfold_thm = Goal.prove_global thy' [] [] ctuple_unfold (fn _ => EVERY [rtac (ctuple_fixdef_thm RS fix_eq2 RS trans) 1, simp_tac (simpset_of thy') 1]); @@ -123,6 +207,17 @@ (*********** monadic notation and pattern matching compilation ***********) (*************************************************************************) +structure FixrecMatchData = TheoryDataFun ( + type T = string Symtab.table; + val empty = Symtab.empty; + val copy = I; + val extend = I; + fun merge _ tabs : T = Symtab.merge (K true) tabs; +); + +(* associate match functions with pattern constants *) +fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms); + fun add_names (Const(a,_), bs) = insert (op =) (Sign.base_name a) bs | add_names (Free(a,_) , bs) = insert (op =) a bs | add_names (f $ u , bs) = add_names (f, add_names(u, bs)) @@ -132,56 +227,63 @@ fun add_terms ts xs = foldr add_names xs ts; (* builds a monadic term for matching a constructor pattern *) -fun pre_build pat rhs vs taken = +fun pre_build match_name pat rhs vs taken = case pat of Const(@{const_name Rep_CFun},_)$f$(v as Free(n,T)) => - pre_build f rhs (v::vs) taken + pre_build match_name f rhs (v::vs) taken | Const(@{const_name Rep_CFun},_)$f$x => - let val (rhs', v, taken') = pre_build x rhs [] taken; - in pre_build f rhs' (v::vs) taken' end + let val (rhs', v, taken') = pre_build match_name x rhs [] taken; + in pre_build match_name f rhs' (v::vs) taken' end | Const(c,T) => let val n = Name.variant taken "v"; fun result_type (Type(@{type_name "->"},[_,T])) (x::xs) = result_type T xs | result_type T _ = T; val v = Free(n, result_type T vs); - val m = "match_"^(extern_name(Sign.base_name c)); + val m = Const(match_name c, matchT T); val k = lambda_ctuple vs rhs; in - (%%:@{const_name Fixrec.bind}`(%%:m`v)`k, v, n::taken) + (mk_bind (m`v, k), v, n::taken) end | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n) | _ => fixrec_err "pre_build: invalid pattern"; (* builds a monadic term for matching a function definition pattern *) (* returns (name, arity, matcher) *) -fun building pat rhs vs taken = +fun building match_name pat rhs vs taken = case pat of Const(@{const_name Rep_CFun}, _)$f$(v as Free(n,T)) => - building f rhs (v::vs) taken + building match_name f rhs (v::vs) taken | Const(@{const_name Rep_CFun}, _)$f$x => - let val (rhs', v, taken') = pre_build x rhs [] taken; - in building f rhs' (v::vs) taken' end - | Const(name,_) => (name, length vs, big_lambdas vs rhs) + let val (rhs', v, taken') = pre_build match_name x rhs [] taken; + in building match_name f rhs' (v::vs) taken' end + | Const(name,_) => (pat, length vs, big_lambdas vs rhs) | _ => fixrec_err "function is not declared as constant in theory"; -fun match_eq eq = +fun match_eq match_name eq = let val (lhs,rhs) = dest_eqs eq; - in building lhs (%%:@{const_name Fixrec.return}`rhs) [] (add_terms [eq] []) end; + in + building match_name lhs (mk_return rhs) [] + (add_terms [eq] []) + end; (* returns the sum (using +++) of the terms in ms *) (* also applies "run" to the result! *) fun fatbar arity ms = let + fun LAM_Ts 0 t = ([], Term.fastype_of t) + | LAM_Ts n (_ $ Abs(_,T,t)) = + let val (Ts, U) = LAM_Ts (n-1) t in (T::Ts, U) end + | LAM_Ts _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; fun unLAM 0 t = t | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; - fun reLAM 0 t = t - | reLAM n t = reLAM (n-1) (%%:@{const_name Abs_CFun} $ Abs("",dummyT,t)); - fun mplus (x,y) = %%:@{const_name Fixrec.mplus}`x`y; - val msum = foldr1 mplus (map (unLAM arity) ms); + fun reLAM ([], U) t = t + | reLAM (T::Ts, U) t = reLAM (Ts, T ->> U) (cabs_const(T,U)$Abs("",T,t)); + val msum = foldr1 mk_mplus (map (unLAM arity) ms); + val (Ts, U) = LAM_Ts arity (hd ms) in - reLAM arity (%%:@{const_name Fixrec.run}`msum) + reLAM (rev Ts, dest_maybeT U) (mk_run msum) end; fun unzip3 [] = ([],[],[]) @@ -190,16 +292,16 @@ in (x::xs, y::ys, z::zs) end; (* this is the pattern-matching compiler function *) -fun compile_pats eqs = +fun compile_pats match_name eqs = let - val ((n::names),(a::arities),mats) = unzip3 (map match_eq eqs); + val ((n::names),(a::arities),mats) = unzip3 (map (match_eq match_name) eqs); val cname = if forall (fn x => n=x) names then n else fixrec_err "all equations in block must define the same function"; val arity = if forall (fn x => a=x) arities then a else fixrec_err "all equations in block must have the same arity"; val rhs = fatbar arity mats; in - mk_trp (%%:cname === rhs) + mk_trp (cname === rhs) end; (*************************************************************************) @@ -235,8 +337,14 @@ fun unconcat [] _ = [] | unconcat (n::ns) xs = List.take (xs,n) :: unconcat ns (List.drop (xs,n)); + val matcher_tab = FixrecMatchData.get thy; + fun match_name c = + case Symtab.lookup matcher_tab c of SOME m => m + | NONE => fixrec_err ("unknown pattern constructor: " ^ c); + val pattern_blocks = unconcat lengths (map Logic.strip_imp_concl eqn_ts'); - val compiled_ts = map (legacy_infer_term thy o compile_pats) pattern_blocks; + val compiled_ts = + map (compile_pats match_name) pattern_blocks; val (thy', cnames, fixdef_thms, unfold_thms) = add_fixdefs compiled_ts thy; in if strict then let (* only prove simp rules if strict = true *) @@ -312,4 +420,6 @@ end; (* local structure *) +val setup = FixrecMatchData.init; + end;