# HG changeset patch # User paulson # Date 1107277317 -3600 # Node ID fc075ae929e433d3e27a31cbd13c770f15db3a07 # Parent cb3612cc41a3b70a9ed8706832bfb2a8c210f910 the new subst tactic, by Lucas Dixon diff -r cb3612cc41a3 -r fc075ae929e4 NEWS --- a/NEWS Sun Jan 30 20:48:50 2005 +0100 +++ b/NEWS Tue Feb 01 18:01:57 2005 +0100 @@ -37,6 +37,12 @@ * Provers/blast.ML: new reference depth_limit to make blast's depth limit (previously hard-coded with a value of 20) user-definable. +* Provers: new version of the subst method, for single-step rewriting: it now + works in bound variable contexts. New is subst (asm), for rewriting an + assumption. Thanks to Lucas Dixon! INCOMPATIBILITY: may rewrite a different + subterm than the original subst method, which is still available under the + name simplesubst. + * Pure: thin_tac now works even if the assumption being deleted contains !! or ==>. More generally, erule now works even if the major premise of the elimination rule contains !! or ==>. diff -r cb3612cc41a3 -r fc075ae929e4 TFL/isand.ML --- a/TFL/isand.ML Sun Jan 30 20:48:50 2005 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,558 +0,0 @@ -(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) -(* Title: sys/isand.ML - Author: Lucas Dixon, University of Edinburgh - lucas.dixon@ed.ac.uk - Date: 6 Aug 2004 -*) -(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) -(* DESCRIPTION: - - Natural Deduction tools - - For working with Isbaelle theorem in a natural detuction style, - ie, not having to deal with meta level quantified varaibles, - instead, we work with newly introduced frees, and hide the - "all"'s, exporting results from theorems proved with the frees, to - solve the all cases of the previous goal. - - Note: A nice idea: allow esxporting to solve any subgoal, thus - allowing the interleaving of proof, or provide a structure for the - ordering of proof, thus allowing proof attempts in parrelle, but - recording the order to apply things in. - -debugging tools: -fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); -fun asm_read s = - (assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); - - THINK: are we really ok with our varify name w.r.t the prop - do - we alos need to avoid named in the hidden hyps? - -*) - -structure IsaND = -struct - -(* Solve *some* subgoal of "th" directly by "sol" *) -(* Note: this is probably what Markus ment to do upon export of a -"show" but maybe he used RS/rtac instead, which would wrongly lead to -failing if there are premices to the shown goal. *) -fun solve_with sol th = - let fun solvei 0 = Seq.empty - | solvei i = - Seq.append (bicompose false (false,sol,0) i th, - solvei (i - 1)) - in - solvei (Thm.nprems_of th) - end; - - -(* Given ctertmify function, (string,type) pairs capturing the free -vars that need to be allified in the assumption, and a theorem with -assumptions possibly containing the free vars, then we give back the -assumptions allified as hidden hyps. *) -(* -Given: vs -th: A vs ==> B vs -Results in: "B vs" [!!x. A x] -*) -fun allify_conditions ctermify Ts th = - let - val premts = Thm.prems_of th - - fun allify_prem_var (vt as (n,ty),t) = - (Term.all ty) $ (Abs(n,ty,Term.abstract_over (Free vt, t))) - - fun allify_prem Ts p = foldr allify_prem_var (Ts, p) - - val cTs = map (ctermify o Free) Ts - val cterm_asms = map (ctermify o allify_prem Ts) premts - val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms - in - (foldl (fn (x,y) => y RS x) (th, allifyied_asm_thms), cterm_asms) - end; - -fun allify_conditions' Ts th = - allify_conditions (Thm.cterm_of (Thm.sign_of_thm th)) Ts th; - - - -(* change schematic vars to fresh free vars *) -fun fix_vars_to_frees th = - let - val ctermify = Thm.cterm_of (Thm.sign_of_thm th) - val prop = Thm.prop_of th - val vars = map Term.dest_Var (Term.term_vars prop) - - val names = Term.add_term_names (prop, []) - - val (insts,names2) = - foldl (fn ((insts,names),v as ((n,i),ty)) => - let val n2 = Term.variant names n in - ((ctermify (Var v), ctermify (Free(n2,ty))) :: insts, - n2 :: names) - end) - (([],names), vars) - in (insts, Thm.instantiate ([], insts) th) end; - -(* *) -(* -val th = Thm.freezeT (topthm()); -val (insts, th') = fix_vars_to_frees th; -val Ts = map (Term.dest_Free o Thm.term_of o snd) insts; -allify_conditions' Ts th'; -*) - - -(* datatype to capture an exported result, ie a fix or assume. *) -datatype export = - export of {fixes : Thm.cterm list, (* fixed vars *) - assumes : Thm.cterm list, (* hidden hyps/assumed prems *) - sgid : int, - gth : Thm.thm}; (* subgoal/goalthm *) - - - -(* export the result of the new goal thm, ie if we reduced teh -subgoal, then we get a new reduced subtgoal with the old -all-quantified variables *) -local -fun allify_term (v, t) = - let val vt = #t (Thm.rep_cterm v) - val (n,ty) = Term.dest_Free vt - in (Term.all ty) $ (Abs(n,ty,Term.abstract_over (vt, t))) end; - -fun allify_for_sg_term ctermify vs t = - let val t_alls = foldr allify_term (vs,t); - val ct_alls = ctermify t_alls; - in - (ct_alls, Drule.forall_elim_list vs (Thm.assume ct_alls)) - end; - -fun lookupfree vs vn = - case Library.find_first (fn (n,ty) => n = vn) vs of - None => raise ERROR_MESSAGE ("prepare_goal_export:lookupfree: " - ^ vn ^ " does not occur in the term") - | Some x => x; -in -fun export_back (export {fixes = vs, assumes = hprems, - sgid = i, gth = gth}) newth = - let - val sgn = Thm.sign_of_thm newth; - val ctermify = Thm.cterm_of sgn; - - val sgs = prems_of newth; - val (sgallcts, sgthms) = - Library.split_list (map (allify_for_sg_term ctermify vs) sgs); - val minimal_newth = - (foldl (fn ( newth', sgthm) => - Drule.compose_single (sgthm, 1, newth')) - (newth, sgthms)); - val allified_newth = - minimal_newth - |> Drule.implies_intr_list hprems - |> Drule.forall_intr_list vs - - val newth' = Drule.implies_intr_list sgallcts allified_newth - in - bicompose false (false, newth', (length sgallcts)) i gth - end; - -(* given a thm of the form: -P1 vs; ...; Pn vs ==> Goal(C vs) - -Gives back: -n, -[| !! vs. P1 vs; !! vs. Pn vs |] - ==> !! C vs -*) -(* note: C may contain further premices etc -Note that cterms is the assumed facts, ie prems of "P1" that are -reintroduced. -*) -fun prepare_goal_export (vs, cterms) th = - let - val sgn = Thm.sign_of_thm th; - val ctermify = Thm.cterm_of sgn; - - val allfrees = map Term.dest_Free (Term.term_frees (Thm.prop_of th)) - val cfrees = map (ctermify o Free o lookupfree allfrees) vs - - val sgs = prems_of th; - val (sgallcts, sgthms) = - Library.split_list (map (allify_for_sg_term ctermify cfrees) sgs); - - val minimal_th = - (foldl (fn ( th', sgthm) => - Drule.compose_single (sgthm, 1, th')) - (th, sgthms)) RS Drule.rev_triv_goal; - - val allified_th = - minimal_th - |> Drule.implies_intr_list cterms - |> Drule.forall_intr_list cfrees - - val th' = Drule.implies_intr_list sgallcts allified_th - in - ((length sgallcts), th') - end; - -end; - -(* val exports_back = foldr (uncurry export_to_sg); *) - -(* test with: - -fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); -fun asm_read s = - (assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); -use_thy "theories/dbg2"; -Goal "!! x :: nat. [| A x; B x; C x; D x |] ==> ((P1 x & P2 x) & P3 x)"; -by (rtac conjI 1); -by (rtac conjI 1); -val th = topthm(); -val i = 1; -val (gthi, exp) = IsaND.fix_alls i th; -val [th'] = Seq.list_of (rtac (thm "p321") 1 gthi); -val oldthewth = Seq.list_of (IsaND.export_back exp th'); - or -val [th'] = Seq.list_of (RWStepTac.rwstep_tac (thms "aa2") 1 gthi); -val oldthewth = Seq.list_of (IsaND.export_back exp th'); - - -val th = topthm(); -val i = 1; -val (goalth, exp1) = IsaND.fix_alls' i th; -val (newth, exp2) = IsaND.hide_other_goals 2 goalth; -val oldthewth = Seq.list_of (IsaND.export_back exp2 newth); -val (export {fixes = vs, assumes = hprems, - sgid = i, gth = gth}) = exp2; - - -Goal "!! x y. P (x + (Suc y)) ==> Z x y ==> Q ((Suc x) + y)"; -val th = topthm(); -val i = 1; -val (gthi, exp) = IsaND.fix_alls i th; -val newth = Seq.hd (Asm_full_simp_tac 1 gthi); -Seq.list_of (IsaND.export_back exp newth); -*) - - - -(* exporting function that takes a solution to the fixed/assumed goal, -and uses this to solve the subgoal in the main theorem *) -fun export_solution (export {fixes = cfvs, assumes = hcprems, - sgid = i, gth = gth}) solth = - let - val solth' = - solth |> Drule.implies_intr_list hcprems - |> Drule.forall_intr_list cfvs - in Drule.compose_single (solth', i, gth) end; - -val exports_solution = foldr (uncurry export_solution); - - - - -(* fix parameters of a subgoal "i", as free variables, and create an -exporting function that will use the result of this proved goal to -show the goal in the original theorem. - -Note, an advantage of this over Isar is that it supports instantiation -of unkowns in the earlier theorem, ie we can do instantiation of meta -vars! *) -(* loosely corresponds to: -Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm -Result: - ("(As ==> SGi x') ==> (As ==> SGi x')" : thm, - expf : - ("As ==> SGi x'" : thm) -> - ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm) -*) -fun fix_alls' i th = - let - val t = (prop_of th); - val names = Term.add_term_names (t,[]); - val gt = Logic.get_goal t i; - val body = Term.strip_all_body gt; - val alls = rev (Term.strip_all_vars gt); - val fvs = map Free - ((Term.variantlist (map fst alls, names)) - ~~ (map snd alls)); - val sgn = Thm.sign_of_thm th; - val ctermify = Thm.cterm_of sgn; - val cfvs = rev (map ctermify fvs); - - val body' = (subst_bounds (fvs,body)); - val gthi0 = Thm.trivial (ctermify body'); - - (* Given a thm justifying subgoal i, solve subgoal i *) - (* Note: fails if there are choices! *) - (* fun exportf thi = - Drule.compose_single (Drule.forall_intr_list cfvs thi, - i, th) *) - in (gthi0, cfvs) end; - -(* small example: -Goal "!! x. [| False; C x |] ==> P x"; -val th = topthm(); -val i = 1; -val (goalth, fixes) = fix_alls' i (topthm()); - -Goal "!! x. P (x + (Suc y)) ==> Q ((Suc x) + y)"; -Goal "!! x. P x y = Q y x ==> P x y"; -val th = topthm(); -val i = 1; -val (prems, gthi, expf) = IsaND.fixes_and_assumes i th; -val gth2 = Seq.hd (RWStepTac.rwstep_tac prems 1 gthi); -*) - - -(* hide other goals *) -(* note the export goal is rotated by (i - 1) and will have to be -unrotated to get backto the originial position(s) *) -fun hide_other_goals th = - let - (* tl beacuse fst sg is the goal we are interested in *) - val cprems = tl (Drule.cprems_of th) - val aprems = map Thm.assume cprems - in - (Drule.implies_elim_list (Drule.rotate_prems 1 th) aprems, - cprems) - end; -(* test with: -Goal "!! x. [| False; C x |] ==> P x"; -val th = topthm(); -val i = 1; -val (goalth, fixedvars) = IsaND.fix_alls' i th; -val (newth, hiddenprems) = IsaND.hide_other_goals goalth; -*) - -(* a nicer version of the above that leaves only a single subgoal (the -other subgoals are hidden hyps, that the exporter suffles about) -namely the subgoal that we were trying to solve. *) -(* loosely corresponds to: -Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm -Result: - ("(As ==> SGi x') ==> SGi x'" : thm, - expf : - ("SGi x'" : thm) -> - ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm) -*) -fun fix_alls i th = - let - val (fixed_gth, fixedvars) = fix_alls' i th - val (sml_gth, othergoals) = hide_other_goals fixed_gth - in - (sml_gth, export {fixes = fixedvars, - assumes = othergoals, - sgid = i, gth = th}) - end; - -(* small example: -Goal "!! x. [| False; C x |] ==> P x"; -val th = topthm(); -val i = 1; -val (goalth, exp) = IsaND.fix_alls i th; -val oldthewth = Seq.list_of (IsaND.export_back exp goalth); - -val (premths, goalth2, expf2) = IsaND.assume_prems 1 goalth; -val False_th = nth_elem (0,premths); -val anything = False_th RS (thm "FalseE"); -val th2 = anything RS goalth2; -val th1 = expf2 th2; -val final_th = flat (map expf th1); -*) - - -(* assume the premises of subgoal "i", this gives back a list of -assumed theorems that are the premices of subgoal i, it also gives -back a new goal thm and an exporter, the new goalthm is as the old -one, but without the premices, and the exporter will use a proof of -the new goalthm, possibly using the assumed premices, to shoe the -orginial goal. *) - -(* Note: Dealing with meta vars, need to meta-level-all them in the -shyps, which we can later instantiate with a specific value.... ? -think about this... maybe need to introduce some new fixed vars and -then remove them again at the end... like I do with rw_inst. *) -(* loosely corresponds to: -Given "[| SG0; ... [| A0; ... An |] ==> SGi; ... SGm |] ==> G" : thm -Result: -(["A0" [A0], ... ,"An" [An]] : thm list, -- assumptions - "SGi ==> SGi" : thm, -- new goal - "SGi" ["A0" ... "An"] : thm -> -- export function - ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list) -*) -fun assume_prems i th = - let - val t = (prop_of th); - val gt = Logic.get_goal t i; - val _ = case Term.strip_all_vars gt of [] => () - | _ => raise ERROR_MESSAGE "assume_prems: goal has params" - val body = gt; - val prems = Logic.strip_imp_prems body; - val concl = Logic.strip_imp_concl body; - - val sgn = Thm.sign_of_thm th; - val ctermify = Thm.cterm_of sgn; - val cprems = map ctermify prems; - val aprems = map Thm.assume cprems; - val gthi = Thm.trivial (ctermify concl); - - (* fun explortf thi = - Drule.compose (Drule.implies_intr_list cprems thi, - i, th) *) - in - (aprems, gthi, cprems) - end; -(* small example: -Goal "False ==> b"; -val th = topthm(); -val i = 1; -val (prems, goalth, cprems) = IsaND.assume_prems i (topthm()); -val F = hd prems; -val FalseE = thm "FalseE"; -val anything = F RS FalseE; -val thi = anything RS goalth; -val res' = expf thi; -*) - - -(* first fix the variables, then assume the assumptions *) -(* loosely corresponds to: -Given - "[| SG0; ... - !! xs. [| A0 xs; ... An xs |] ==> SGi xs; - ... SGm |] ==> G" : thm -Result: -(["A0 xs'" [A0 xs'], ... ,"An xs'" [An xs']] : thm list, -- assumptions - "SGi xs' ==> SGi xs'" : thm, -- new goal - "SGi xs'" ["A0 xs'" ... "An xs'"] : thm -> -- export function - ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list) -*) - -(* Note: the fix_alls actually pulls through all the assumptions which -means that the second export is not needed. *) -fun fixes_and_assumes i th = - let - val (fixgth, exp1) = fix_alls i th - val (assumps, goalth, _) = assume_prems 1 fixgth - in - (assumps, goalth, exp1) - end; -(* test with: -Goal "!! x. [| False; C x |] ==> P x"; -val th = topthm(); -val i = 1; -val (fixgth, exp) = IsaND.fix_alls i th; -val (assumps, goalth, _) = assume_prems 1 fixgth; - -val oldthewth = Seq.list_of (IsaND.export_back exp fixgth); -val oldthewth = Seq.list_of (IsaND.export_back exp goalth); - - -val (prems, goalth, expf) = IsaND.fixes_and_assumes i th; -val F = hd prems; -val FalseE = thm "FalseE"; -val anything = F RS FalseE; -val thi = anything RS goalth; -val res' = expf thi; -*) - -(* Fixme: allow different order of subgoals given to expf *) -(* make each subgoal into a separate thm that needs to be proved *) -(* loosely corresponds to: -Given - "[| SG0; ... SGm |] ==> G" : thm -Result: -(["SG0 ==> SG0", ... ,"SGm ==> SGm"] : thm list, -- goals - ["SG0", ..., "SGm"] : thm list -> -- export function - "G" : thm) -*) -fun subgoal_thms th = - let - val t = (prop_of th); - - val prems = Logic.strip_imp_prems t; - - val sgn = Thm.sign_of_thm th; - val ctermify = Thm.cterm_of sgn; - - val aprems = map (Thm.trivial o ctermify) prems; - - fun explortf premths = - Drule.implies_elim_list th premths - in - (aprems, explortf) - end; -(* small example: -Goal "A & B"; -by (rtac conjI 1); -val th = topthm(); -val (subgoals, expf) = subgoal_thms (topthm()); -*) - -(* make all the premices of a theorem hidden, and provide an unhide -function, that will bring them back out at a later point. This is -useful if you want to get back these premices, after having used the -theorem with the premices hidden *) -(* loosely corresponds to: -Given "As ==> G" : thm -Result: ("G [As]" : thm, - "G [As]" : thm -> "As ==> G" : thm -*) -fun hide_prems th = - let - val sgn = Thm.sign_of_thm th; - val ctermify = Thm.cterm_of sgn; - - val t = (prop_of th); - val prems = Logic.strip_imp_prems t; - val cprems = map ctermify prems; - val aprems = map Thm.trivial cprems; - - (* val unhidef = Drule.implies_intr_list cprems; *) - in - (Drule.implies_elim_list th aprems, cprems) - end; - - - - -(* Fixme: allow different order of subgoals in exportf *) -(* as above, but also fix all parameters in all subgoals, and uses -fix_alls, not fix_alls', ie doesn't leave extra asumptions as apparent -subgoals. *) -(* loosely corresponds to: -Given - "[| !! x0s. A0s x0s ==> SG0 x0s; - ...; !! xms. Ams xms ==> SGm xms|] ==> G" : thm -Result: -(["(A0s x0s' ==> SG0 x0s') ==> SG0 x0s'", - ... ,"(Ams xms' ==> SGm xms') ==> SGm xms'"] : thm list, -- goals - ["SG0 x0s'", ..., "SGm xms'"] : thm list -> -- export function - "G" : thm) -*) -(* requires being given solutions! *) -fun fixed_subgoal_thms th = - let - val (subgoals, expf) = subgoal_thms th; -(* fun export_sg (th, exp) = exp th; *) - fun export_sgs expfs solthms = - expf (map2 (op |>) (solthms, expfs)); -(* expf (map export_sg (ths ~~ expfs)); *) - in - apsnd export_sgs (Library.split_list (map (apsnd export_solution o - fix_alls 1) subgoals)) - end; - - -(* small example: -Goal "(!! x. ((C x) ==> (A x)))"; -val th = topthm(); -val i = 1; -val (subgoals, expf) = fixed_subgoal_thms (topthm()); -*) - -end; diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Advanced/document/Partial.tex --- a/doc-src/TutorialI/Advanced/document/Partial.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Advanced/document/Partial.tex Tue Feb 01 18:01:57 2005 +0100 @@ -150,7 +150,7 @@ \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymLongrightarrow}\ find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ f\ x\ {\isacharequal}\ x\ then\ x\ else\ find{\isacharparenleft}f{\isacharcomma}\ f\ x{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent Then you should disable the original recursion equation:% @@ -165,11 +165,9 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymlongrightarrow}\ f{\isacharparenleft}find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f\ x\ rule{\isacharcolon}\ find{\isachardot}induct{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \isamarkupsubsubsection{The {\tt\slshape while} Combinator% } @@ -231,14 +229,10 @@ \ \ {\isasymexists}y{\isachardot}\ while\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ {\isacharparenleft}x{\isacharprime}{\isacharcomma}f\ x{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}x{\isacharcomma}f\ x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}y{\isacharparenright}\ {\isasymand}\isanewline \ \ \ \ \ \ \ f\ y\ {\isacharequal}\ y{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ P\ {\isacharequal}\ {\isachardoublequote}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isacharequal}\ f\ x{\isachardoublequote}\ \isakeyword{and}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ r\ {\isacharequal}\ {\isachardoublequote}inv{\isacharunderscore}image\ {\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ fst{\isachardoublequote}\ \isakeyword{in}\ while{\isacharunderscore}rule{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ inv{\isacharunderscore}image{\isacharunderscore}def\ step{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% The theorem itself is a simple consequence of this lemma:% @@ -246,11 +240,9 @@ \isamarkuptrue% \isacommand{theorem}\ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymLongrightarrow}\ f{\isacharparenleft}find{\isadigit{2}}\ f\ x{\isacharparenright}\ {\isacharequal}\ find{\isadigit{2}}\ f\ x{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}drule{\isacharunderscore}tac\ x\ {\isacharequal}\ x\ \isakeyword{in}\ lem{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ find{\isadigit{2}}{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% Let us conclude this section on partial functions by a diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Advanced/document/WFrec.tex --- a/doc-src/TutorialI/Advanced/document/WFrec.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Advanced/document/WFrec.tex Tue Feb 01 18:01:57 2005 +0100 @@ -85,37 +85,12 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ wf{\isacharunderscore}greater{\isacharcolon}\ {\isachardoublequote}wf\ {\isacharbraceleft}{\isacharparenleft}i{\isacharcomma}j{\isacharparenright}{\isachardot}\ j{\isacharless}i\ {\isasymand}\ i\ {\isasymle}\ {\isacharparenleft}N{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The proof is by showing that our relation is a subset of another well-founded -relation: one given by a measure function.\index{*wf_subset (theorem)}% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}rule\ wf{\isacharunderscore}subset\ {\isacharbrackleft}of\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}k{\isacharcolon}{\isacharcolon}nat{\isachardot}\ N{\isacharminus}k{\isacharparenright}{\isachardoublequote}{\isacharbrackright}{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isacharbraceleft}{\isacharparenleft}i{\isacharcomma}\ j{\isacharparenright}{\isachardot}\ j\ {\isacharless}\ i\ {\isasymand}\ i\ {\isasymle}\ N{\isacharbraceright}\ {\isasymsubseteq}\ measure\ {\isacharparenleft}op\ {\isacharminus}\ N{\isacharparenright}% -\end{isabelle} - -\noindent -The inclusion remains to be proved. After unfolding some definitions, -we are left with simple arithmetic:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}clarify{\isacharcomma}\ simp\ add{\isacharcolon}\ measure{\isacharunderscore}def\ inv{\isacharunderscore}image{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ b{\isachardot}\ {\isasymlbrakk}b\ {\isacharless}\ a{\isacharsemicolon}\ a\ {\isasymle}\ N{\isasymrbrakk}\ {\isasymLongrightarrow}\ N\ {\isacharminus}\ a\ {\isacharless}\ N\ {\isacharminus}\ b% -\end{isabelle} - -\noindent -And that is dispatched automatically:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{by}\ arith\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/CodeGen/document/CodeGen.tex --- a/doc-src/TutorialI/CodeGen/document/CodeGen.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/CodeGen/document/CodeGen.tex Tue Feb 01 18:01:57 2005 +0100 @@ -105,27 +105,12 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{theorem}\ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}comp\ e{\isacharparenright}\ s\ vs\ {\isacharequal}\ {\isacharparenleft}value\ e\ s{\isacharparenright}\ {\isacharhash}\ vs{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -It will be proved by induction on \isa{e} followed by simplification. -First, we must prove a lemma about executing the concatenation of two -instruction sequences:% -\end{isamarkuptxt}% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ exec{\isacharunderscore}app{\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}xs{\isacharat}ys{\isacharparenright}\ s\ vs\ {\isacharequal}\ exec\ ys\ s\ {\isacharparenleft}exec\ xs\ s\ vs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This requires induction on \isa{xs} and ordinary simplification for the -base cases. In the induction step, simplification leaves us with a formula -that contains two \isa{case}-expressions over instructions. Thus we add -automatic case splitting, which finishes the proof:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -137,7 +122,7 @@ \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Datatype/document/ABexpr.tex --- a/doc-src/TutorialI/Datatype/document/ABexpr.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Datatype/document/ABexpr.tex Tue Feb 01 18:01:57 2005 +0100 @@ -103,14 +103,9 @@ \isacommand{lemma}\ {\isachardoublequote}evala\ {\isacharparenleft}substa\ s\ a{\isacharparenright}\ env\ {\isacharequal}\ evala\ a\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ evala\ {\isacharparenleft}s\ x{\isacharparenright}\ env{\isacharparenright}\ {\isasymand}\isanewline \ \ \ \ \ \ \ \ evalb\ {\isacharparenleft}substb\ s\ b{\isacharparenright}\ env\ {\isacharequal}\ evalb\ b\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ evala\ {\isacharparenleft}s\ x{\isacharparenright}\ env{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ a\ \isakeyword{and}\ b{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The resulting 8 goals (one for each constructor) are proved in one fell swoop:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ simp{\isacharunderscore}all\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Datatype/document/Fundata.tex --- a/doc-src/TutorialI/Datatype/document/Fundata.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Datatype/document/Fundata.tex Tue Feb 01 18:01:57 2005 +0100 @@ -43,22 +43,10 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}map{\isacharunderscore}bt\ {\isacharparenleft}g\ o\ f{\isacharparenright}\ T\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ T{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ T{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{done}\isamarkupfalse% \isamarkupfalse% \isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Because of the function type, the proof state after induction looks unusual. -Notice the quantified induction hypothesis: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ map{\isacharunderscore}bt\ {\isacharparenleft}g\ {\isasymcirc}\ f{\isacharparenright}\ Tip\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ Tip{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ F{\isachardot}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ map{\isacharunderscore}bt\ {\isacharparenleft}g\ {\isasymcirc}\ f{\isacharparenright}\ {\isacharparenleft}F\ x{\isacharparenright}\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ {\isacharparenleft}F\ x{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ F{\isachardot}\ }map{\isacharunderscore}bt\ {\isacharparenleft}g\ {\isasymcirc}\ f{\isacharparenright}\ {\isacharparenleft}Br\ a\ F{\isacharparenright}\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ {\isacharparenleft}Br\ a\ F{\isacharparenright}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Datatype/document/Nested.tex --- a/doc-src/TutorialI/Datatype/document/Nested.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Datatype/document/Nested.tex Tue Feb 01 18:01:57 2005 +0100 @@ -73,9 +73,8 @@ \isacommand{lemma}\ subst{\isacharunderscore}id{\isacharcolon}\ {\isachardoublequote}subst\ \ Var\ t\ \ {\isacharequal}\ {\isacharparenleft}t\ {\isacharcolon}{\isacharcolon}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharparenright}\ \ {\isasymand}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ substs\ Var\ ts\ {\isacharequal}\ {\isacharparenleft}ts{\isacharcolon}{\isacharcolon}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ \isakeyword{and}\ ts{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -129,9 +128,8 @@ \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}subst\ s\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}map\ {\isacharparenleft}subst\ s{\isacharparenright}\ ts{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Documents/document/Documents.tex --- a/doc-src/TutorialI/Documents/document/Documents.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Documents/document/Documents.tex Tue Feb 01 18:01:57 2005 +0100 @@ -587,9 +587,6 @@ } \isanewline \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}rule\ impI{\isacharparenright}\ % -\isamarkupcmt{implicit assumption step involved here% -} \isamarkupfalse% % \begin{isamarkuptext}% @@ -815,7 +812,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}x\ {\isasymnoteq}\ {\isacharparenleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ {\isasymLongrightarrow}\ {\isadigit{0}}\ {\isacharless}\ x\ {\isacharasterisk}\ x{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent Here the real source of the proof has been as follows: diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Ifexpr/document/Ifexpr.tex --- a/doc-src/TutorialI/Ifexpr/document/Ifexpr.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Ifexpr/document/Ifexpr.tex Tue Feb 01 18:01:57 2005 +0100 @@ -100,17 +100,10 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}valif\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ env\ {\isacharequal}\ value\ b\ env{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The proof is canonical:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ b{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Inductive/document/AB.tex --- a/doc-src/TutorialI/Inductive/document/AB.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Inductive/document/AB.tex Tue Feb 01 18:01:57 2005 +0100 @@ -38,7 +38,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ b{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}x\ {\isasymnoteq}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ a{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}case{\isacharunderscore}tac\ x{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -80,16 +80,8 @@ \ \ {\isachardoublequote}{\isacharparenleft}w\ {\isasymin}\ S\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}{\isacharparenright}\ \ \ \ \ {\isasymand}\isanewline \ \ \ {\isacharparenleft}w\ {\isasymin}\ A\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymand}\isanewline \ \ \ {\isacharparenleft}w\ {\isasymin}\ B\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -These propositions are expressed with the help of the predefined \isa{filter} function on lists, which has the convenient syntax \isa{{\isacharbrackleft}x{\isasymin}xs{\isachardot}\ P\ x{\isacharbrackright}}, the list of all elements \isa{x} in \isa{xs} such that \isa{P\ x} -holds. Remember that on lists \isa{size} and \isa{length} are synonymous. - -The proof itself is by rule induction and afterwards automatic:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{by}\ {\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -128,26 +120,10 @@ \isacommand{lemma}\ step{\isadigit{1}}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i\ {\isacharless}\ size\ w{\isachardot}\isanewline \ \ {\isasymbar}{\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\isanewline \ \ \ {\isacharminus}\ {\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isasymbar}\ {\isasymle}\ {\isadigit{1}}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The lemma is a bit hard to read because of the coercion function -\isa{int\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ int}. It is required because \isa{size} returns -a natural number, but subtraction on type~\isa{nat} will do the wrong thing. -Function \isa{take} is predefined and \isa{take\ i\ xs} is the prefix of -length \isa{i} of \isa{xs}; below we also need \isa{drop\ i\ xs}, which -is what remains after that prefix has been dropped from \isa{xs}. - -The proof is by induction on \isa{w}, with a trivial base case, and a not -so trivial induction step. Since it is essentially just arithmetic, we do not -discuss it.% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ zabs{\isacharunderscore}def\ take{\isacharunderscore}Cons\ split{\isacharcolon}\ nat{\isachardot}split\ if{\isacharunderscore}splits{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Finally we come to the above-mentioned lemma about cutting in half a word with two more elements of one sort than of the other sort:% @@ -156,17 +132,9 @@ \isacommand{lemma}\ part{\isadigit{1}}{\isacharcolon}\isanewline \ {\isachardoublequote}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{2}}\ {\isasymLongrightarrow}\isanewline \ \ {\isasymexists}i{\isasymle}size\ w{\isachardot}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This is proved by \isa{force} with the help of the intermediate value theorem, -instantiated appropriately and with its first premise disposed of by lemma -\isa{step{\isadigit{1}}}:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}insert\ nat{\isadigit{0}}{\isacharunderscore}intermed{\isacharunderscore}int{\isacharunderscore}val{\isacharbrackleft}OF\ step{\isadigit{1}}{\isacharcomma}\ of\ {\isachardoublequote}P{\isachardoublequote}\ {\isachardoublequote}w{\isachardoublequote}\ {\isachardoublequote}{\isadigit{1}}{\isachardoublequote}{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}\ force\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -181,7 +149,7 @@ \ \ \ \ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isasymrbrakk}\isanewline \ \ \ {\isasymLongrightarrow}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ del{\isacharcolon}\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -214,119 +182,38 @@ \ \ {\isachardoublequote}{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ \ \ \ \ {\isasymlongrightarrow}\ w\ {\isasymin}\ S{\isacharparenright}\ {\isasymand}\isanewline \ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ A{\isacharparenright}\ {\isasymand}\isanewline \ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The proof is by induction on \isa{w}. Structural induction would fail here -because, as we can see from the grammar, we need to make bigger steps than -merely appending a single letter at the front. Hence we induct on the length -of \isa{w}, using the induction rule \isa{length{\isacharunderscore}induct}:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w\ rule{\isacharcolon}\ length{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The \isa{rule} parameter tells \isa{induct{\isacharunderscore}tac} explicitly which induction -rule to use. For details see \S\ref{sec:complete-ind} below. -In this case the result is that we may assume the lemma already -holds for all words shorter than \isa{w}. - -The proof continues with a case distinction on \isa{w}, -on whether \isa{w} is empty or not.% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ w{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Simplification disposes of the base case and leaves only a conjunction -of two step cases to be proved: -if \isa{w\ {\isacharequal}\ a\ {\isacharhash}\ v} and \begin{isabelle}% -\ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{2}}% -\end{isabelle} then -\isa{b\ {\isacharhash}\ v\ {\isasymin}\ A}, and similarly for \isa{w\ {\isacharequal}\ b\ {\isacharhash}\ v}. -We only consider the first case in detail. - -After breaking the conjunction up into two cases, we can apply -\isa{part{\isadigit{1}}} to the assumption that \isa{w} contains two more \isa{a}'s than \isa{b}'s.% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This yields an index \isa{i\ {\isasymle}\ length\ v} such that -\begin{isabelle}% -\ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}take\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}take\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}% -\end{isabelle} -With the help of \isa{part{\isadigit{2}}} it follows that -\begin{isabelle}% -\ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}% -\end{isabelle}% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}drule\ part{\isadigit{2}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Now it is time to decompose \isa{v} in the conclusion \isa{b\ {\isacharhash}\ v\ {\isasymin}\ A} -into \isa{take\ i\ v\ {\isacharat}\ drop\ i\ v},% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -(the variables \isa{n{\isadigit{1}}} and \isa{t} are the result of composing the -theorems \isa{subst} and \isa{append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id}) -after which the appropriate rule of the grammar reduces the goal -to the two subgoals \isa{take\ i\ v\ {\isasymin}\ A} and \isa{drop\ i\ v\ {\isasymin}\ A}:% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Both subgoals follow from the induction hypothesis because both \isa{take\ i\ v} and \isa{drop\ i\ v} are shorter than \isa{w}:% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}force\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -The case \isa{w\ {\isacharequal}\ b\ {\isacharhash}\ v} is proved analogously:% -\end{isamarkuptxt}% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}b{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}drule\ part{\isadigit{2}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}b{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% We conclude this section with a comparison of our proof with diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Inductive/document/Advanced.tex --- a/doc-src/TutorialI/Inductive/document/Advanced.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Inductive/document/Advanced.tex Tue Feb 01 18:01:57 2005 +0100 @@ -22,21 +22,11 @@ \isamarkupfalse% \isacommand{lemma}\ gterms{\isacharunderscore}mono{\isacharcolon}\ {\isachardoublequote}F{\isasymsubseteq}G\ {\isasymLongrightarrow}\ gterms\ F\ {\isasymsubseteq}\ gterms\ G{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ clarify\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}F\ {\isasymsubseteq}\ G{\isacharsemicolon}\ {\isasymforall}t{\isasymin}set\ args{\isachardot}\ t\ {\isasymin}\ gterms\ F\ {\isasymand}\ t\ {\isasymin}\ gterms\ G{\isacharsemicolon}\ f\ {\isasymin}\ F{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ gterms\ G% -\end{isabelle}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -67,9 +57,8 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}Suc{\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ P\ n{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}ind{\isacharunderscore}cases\ {\isachardoublequote}Suc{\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even{\isachardoublequote}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{oops}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{inductive{\isacharunderscore}cases}\ gterm{\isacharunderscore}Apply{\isacharunderscore}elim\ {\isacharbrackleft}elim{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}Apply\ f\ args\ {\isasymin}\ gterms\ F{\isachardoublequote}\isamarkupfalse% @@ -86,22 +75,10 @@ \isacommand{lemma}\ gterms{\isacharunderscore}IntI\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\isanewline \ \ \ \ \ {\isachardoublequote}t\ {\isasymin}\ gterms\ F\ {\isasymLongrightarrow}\ t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}args\ f{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}\ \ \ }t\ {\isasymin}\ gterms\ F\ {\isasymand}\ {\isacharparenleft}t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }f\ {\isasymin}\ F{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymLongrightarrow}\ }Apply\ f\ args\ {\isasymin}\ gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -113,7 +90,7 @@ \isacommand{lemma}\ gterms{\isacharunderscore}Int{\isacharunderscore}eq\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline \ \ \ \ \ {\isachardoublequote}gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}\ {\isacharequal}\ gterms\ F\ {\isasyminter}\ gterms\ G{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ mono{\isacharunderscore}Int\ monoI\ gterms{\isacharunderscore}mono{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% the following declaration isn't actually used% @@ -149,67 +126,24 @@ \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ clarify\isamarkupfalse% -% -\begin{isamarkuptxt}% -The situation after clarify -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ }x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -note the induction hypothesis! -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}\ \ \ }t\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymand}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}\ \ \ }t\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isanewline \isanewline \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ clarify\isamarkupfalse% -% -\begin{isamarkuptxt}% -The situation after clarify -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ }x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -note the induction hypothesis! -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}args\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}}{\isasymin}\ lists\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}{\isasymin}\ \ }{\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasyminter}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}{\isasymin}\ \ {\isacharparenleft}}{\isacharbraceleft}x{\isachardot}\ x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isacharbraceright}{\isacharparenright}{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -258,24 +192,18 @@ \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}typed{\isacharunderscore}gterm\ sig\ {\isasymsubseteq}\ well{\isacharunderscore}typed{\isacharunderscore}gterm{\isacharprime}\ sig{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ clarify\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}typed{\isacharunderscore}gterm{\isachardot}induct{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}typed{\isacharunderscore}gterm{\isacharprime}\ sig\ {\isasymsubseteq}\ well{\isacharunderscore}typed{\isacharunderscore}gterm\ sig{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ clarify\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}typed{\isacharunderscore}gterm{\isacharprime}{\isachardot}induct{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isanewline \isanewline \isanewline \isamarkupfalse% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Inductive/document/Even.tex --- a/doc-src/TutorialI/Inductive/document/Even.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Inductive/document/Even.tex Tue Feb 01 18:01:57 2005 +0100 @@ -35,22 +35,10 @@ \isamarkuptrue% \isacommand{lemma}\ two{\isacharunderscore}times{\isacharunderscore}even{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isadigit{2}}{\isacharasterisk}k\ {\isasymin}\ even{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}induct{\isacharunderscore}tac\ k{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -The first step is induction on the natural number \isa{k}, which leaves -two subgoals: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{2}}\ {\isacharasterisk}\ {\isadigit{0}}\ {\isasymin}\ Even{\isachardot}even\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{2}}\ {\isacharasterisk}\ n\ {\isasymin}\ Even{\isachardot}even\ {\isasymLongrightarrow}\ {\isadigit{2}}\ {\isacharasterisk}\ Suc\ n\ {\isasymin}\ Even{\isachardot}even% -\end{isabelle} -Here \isa{auto} simplifies both subgoals so that they match the introduction -rules, which then are applied automatically.% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Our goal is to prove the equivalence between the traditional definition @@ -61,7 +49,7 @@ \isamarkuptrue% \isacommand{lemma}\ dvd{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}{\isadigit{2}}\ dvd\ n\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% our first rule induction!% @@ -69,34 +57,14 @@ \isamarkuptrue% \isacommand{lemma}\ even{\isacharunderscore}imp{\isacharunderscore}dvd{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ n{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{2}}\ dvd\ {\isadigit{0}}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ {\isadigit{2}}\ dvd\ n{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ {\isasymexists}k{\isachardot}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ k{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymexists}k{\isachardot}\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ k% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ clarify\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n\ k{\isachardot}\ {\isadigit{2}}\ {\isacharasterisk}\ k\ {\isasymin}\ Even{\isachardot}even\ {\isasymLongrightarrow}\ {\isasymexists}ka{\isachardot}\ Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ k{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ ka% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequote}Suc\ k{\isachardoublequote}\ \isakeyword{in}\ exI{\isacharcomma}\ simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% no iff-attribute because we don't always want to use it% @@ -104,7 +72,7 @@ \isamarkuptrue% \isacommand{theorem}\ even{\isacharunderscore}iff{\isacharunderscore}dvd{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}n\ {\isasymin}\ even{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{2}}\ dvd\ n{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharcolon}\ dvd{\isacharunderscore}imp{\isacharunderscore}even\ even{\isacharunderscore}imp{\isacharunderscore}dvd{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% this result ISN'T inductive...% @@ -112,16 +80,9 @@ \isamarkuptrue% \isacommand{lemma}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ n\ {\isasymin}\ Even{\isachardot}even\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}na{\isachardot}\ {\isasymlbrakk}na\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ n\ {\isasymin}\ Even{\isachardot}even{\isasymrbrakk}\ {\isasymLongrightarrow}\ n\ {\isasymin}\ Even{\isachardot}even% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{oops}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% ...so we need an inductive lemma...% @@ -129,19 +90,10 @@ \isamarkuptrue% \isacommand{lemma}\ even{\isacharunderscore}imp{\isacharunderscore}even{\isacharunderscore}minus{\isacharunderscore}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ even{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{0}}\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ Even{\isachardot}even\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ n\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ Even{\isachardot}even{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ }{\isasymLongrightarrow}\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ Even{\isachardot}even% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% ...and prove it in a separate step% @@ -149,13 +101,13 @@ \isamarkuptrue% \isacommand{lemma}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}drule\ even{\isacharunderscore}imp{\isacharunderscore}even{\isacharunderscore}minus{\isacharunderscore}{\isadigit{2}}{\isacharcomma}\ simp{\isacharparenright}\isanewline +\isanewline \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymin}\ even{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymin}\ even{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ dest{\isacharcolon}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharparenright}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{end}\isanewline diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Inductive/document/Mutual.tex --- a/doc-src/TutorialI/Inductive/document/Mutual.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Inductive/document/Mutual.tex Tue Feb 01 18:01:57 2005 +0100 @@ -39,27 +39,8 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}m\ {\isasymin}\ even\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ m{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}n\ {\isasymin}\ odd\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The proof is by rule induction. Because of the form of the induction theorem, -it is applied by \isa{rule} rather than \isa{erule} as for ordinary -inductive definitions:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{2}}\ dvd\ {\isadigit{0}}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ odd{\isacharsemicolon}\ {\isadigit{2}}\ dvd\ Suc\ n{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ Suc\ n\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Mutual{\isachardot}even{\isacharsemicolon}\ {\isadigit{2}}\ dvd\ n{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}% -\end{isabelle} -The first two subgoals are proved by simplification and the final one can be -proved in the same manner as in \S\ref{sec:rule-induction} -where the same subgoal was encountered before. -We do not show the proof script.% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Inductive/document/Star.tex --- a/doc-src/TutorialI/Inductive/document/Star.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Inductive/document/Star.tex Tue Feb 01 18:01:57 2005 +0100 @@ -47,7 +47,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -76,69 +76,17 @@ \isamarkuptrue% \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Unfortunately, even the base case is a problem: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}% -\end{isabelle} -We have to abandon this proof attempt. -To understand what is going on, let us look again at \isa{rtc{\isachardot}induct}. -In the above application of \isa{erule}, the first premise of -\isa{rtc{\isachardot}induct} is unified with the first suitable assumption, which -is \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}} rather than \isa{{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}}. Although that -is what we want, it is merely due to the order in which the assumptions occur -in the subgoal, which it is not good practice to rely on. As a result, -\isa{{\isacharquery}xb} becomes \isa{x}, \isa{{\isacharquery}xa} becomes -\isa{y} and \isa{{\isacharquery}P} becomes \isa{{\isasymlambda}u\ v{\isachardot}\ {\isacharparenleft}u{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}}, thus -yielding the above subgoal. So what went wrong? - -When looking at the instantiation of \isa{{\isacharquery}P} we see that it does not -depend on its second parameter at all. The reason is that in our original -goal, of the pair \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}} only \isa{x} appears also in the -conclusion, but not \isa{y}. Thus our induction statement is too -weak. Fortunately, it can easily be strengthened: -transfer the additional premise \isa{{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}} into the conclusion:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This is not an obscure trick but a generally applicable heuristic: -\begin{quote}\em -When proving a statement by rule induction on $(x@1,\dots,x@n) \in R$, -pull all other premises containing any of the $x@i$ into the conclusion -using $\longrightarrow$. -\end{quote} -A similar heuristic for other kinds of inductions is formulated in -\S\ref{sec:ind-var-in-prems}. The \isa{rule{\isacharunderscore}format} directive turns -\isa{{\isasymlongrightarrow}} back into \isa{{\isasymLongrightarrow}}: in the end we obtain the original -statement of our lemma.% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Now induction produces two subgoals which are both proved automatically: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y\ za{\isachardot}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}\ za{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}% -\end{isabelle}% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% Let us now prove that \isa{r{\isacharasterisk}} is really the reflexive transitive closure @@ -162,26 +110,19 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isadigit{2}}{\isachardot}induct{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline +\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}trans{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isadigit{2}}{\isachardot}intros{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isadigit{2}}{\isachardot}intros{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% So why did we start with the first definition? Because it is simpler. It diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Misc/document/AdvancedInd.tex --- a/doc-src/TutorialI/Misc/document/AdvancedInd.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Misc/document/AdvancedInd.tex Tue Feb 01 18:01:57 2005 +0100 @@ -29,99 +29,11 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -But induction produces the warning -\begin{quote}\tt -Induction variable occurs also among premises! -\end{quote} -and leads to the base case -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd\ {\isacharparenleft}rev\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ last\ {\isacharbrackleft}{\isacharbrackright}% -\end{isabelle} -Simplification reduces the base case to this: -\begin{isabelle} -\ 1.\ xs\ {\isasymnoteq}\ []\ {\isasymLongrightarrow}\ hd\ []\ =\ last\ [] -\end{isabelle} -We cannot prove this equality because we do not know what \isa{hd} and -\isa{last} return when applied to \isa{{\isacharbrackleft}{\isacharbrackright}}. - -We should not have ignored the warning. Because the induction -formula is only the conclusion, induction does not affect the occurrence of \isa{xs} in the premises. -Thus the case that should have been trivial -becomes unprovable. Fortunately, the solution is easy:\footnote{A similar -heuristic applies to rule inductions; see \S\ref{sec:rtc}.} -\begin{quote} -\emph{Pull all occurrences of the induction variable into the conclusion -using \isa{{\isasymlongrightarrow}}.} -\end{quote} -Thus we should state the lemma as an ordinary -implication~(\isa{{\isasymlongrightarrow}}), letting -\attrdx{rule_format} (\S\ref{sec:forward}) convert the -result to the usual \isa{{\isasymLongrightarrow}} form:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ hd{\isacharunderscore}rev\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This time, induction leaves us with a trivial base case: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd\ {\isacharparenleft}rev\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ last\ {\isacharbrackleft}{\isacharbrackright}% -\end{isabelle} -And \isa{auto} completes the proof. - -If there are multiple premises $A@1$, \dots, $A@n$ containing the -induction variable, you should turn the conclusion $C$ into -\[ A@1 \longrightarrow \cdots A@n \longrightarrow C. \] -Additionally, you may also have to universally quantify some other variables, -which can yield a fairly complex conclusion. However, \isa{rule{\isacharunderscore}format} -can remove any number of occurrences of \isa{{\isasymforall}} and -\isa{{\isasymlongrightarrow}}. - -\index{induction!on a term}% -A second reason why your proposition may not be amenable to induction is that -you want to induct on a complex term, rather than a variable. In -general, induction on a term~$t$ requires rephrasing the conclusion~$C$ -as -\begin{equation}\label{eqn:ind-over-term} -\forall y@1 \dots y@n.~ x = t \longrightarrow C. -\end{equation} -where $y@1 \dots y@n$ are the free variables in $t$ and $x$ is a new variable. -Now you can perform induction on~$x$. An example appears in -\S\ref{sec:complete-ind} below. - -The very same problem may occur in connection with rule induction. Remember -that it requires a premise of the form $(x@1,\dots,x@k) \in R$, where $R$ is -some inductively defined set and the $x@i$ are variables. If instead we have -a premise $t \in R$, where $t$ is not just an $n$-tuple of variables, we -replace it with $(x@1,\dots,x@k) \in R$, and rephrase the conclusion $C$ as -\[ \forall y@1 \dots y@n.~ (x@1,\dots,x@k) = t \longrightarrow C. \] -For an example see \S\ref{sec:CTL-revisited} below. - -Of course, all premises that share free variables with $t$ need to be pulled into -the conclusion as well, under the \isa{{\isasymforall}}, again using \isa{{\isasymlongrightarrow}} as shown above. - -Readers who are puzzled by the form of statement -(\ref{eqn:ind-over-term}) above should remember that the -transformation is only performed to permit induction. Once induction -has been applied, the statement can be transformed back into something quite -intuitive. For example, applying wellfounded induction on $x$ (w.r.t.\ -$\prec$) to (\ref{eqn:ind-over-term}) and transforming the result a -little leads to the goal -\[ \bigwedge\overline{y}.\ - \forall \overline{z}.\ t\,\overline{z} \prec t\,\overline{y}\ \longrightarrow\ C\,\overline{z} - \ \Longrightarrow\ C\,\overline{y} \] -where $\overline{y}$ stands for $y@1 \dots y@n$ and the dependence of $t$ and -$C$ on the free variables of $t$ has been made explicit. -Unfortunately, this induction schema cannot be expressed as a -single theorem because it depends on the number of free variables in $t$ --- -the notation $\overline{y}$ is merely an informal device.% -\end{isamarkuptxt}% \isamarkuptrue% \isamarkupfalse% % @@ -170,41 +82,14 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -To perform induction on \isa{k} using \isa{nat{\isacharunderscore}less{\isacharunderscore}induct}, we use -the same general induction method as for recursion induction (see -\S\ref{sec:recdef-induction}):% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ k\ rule{\isacharcolon}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -We get the following proof state: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ {\isasymforall}i{\isachardot}\ m\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i\ {\isasymLongrightarrow}\ {\isasymforall}i{\isachardot}\ n\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i% -\end{isabelle} -After stripping the \isa{{\isasymforall}i}, the proof continues with a case -distinction on \isa{i}. The case \isa{i\ {\isacharequal}\ {\isadigit{0}}} is trivial and we focus on -the other case:% -\end{isamarkuptxt}% +\isamarkupfalse% +\isamarkupfalse% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ allI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ i{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n\ i\ nat{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}m{\isacharless}n{\isachardot}\ {\isasymforall}i{\isachardot}\ m\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isacharsemicolon}\ i\ {\isacharequal}\ Suc\ nat{\isasymrbrakk}\ {\isasymLongrightarrow}\ n\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{by}{\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ f{\isacharunderscore}ax\ Suc{\isacharunderscore}leI\ intro{\isacharcolon}\ le{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -295,18 +180,10 @@ \isamarkuptrue% \isacommand{lemma}\ induct{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The base case is vacuously true. For the induction step (\isa{m\ {\isacharless}\ Suc\ n}) we distinguish two cases: case \isa{m\ {\isacharless}\ n} is true by induction -hypothesis and case \isa{m\ {\isacharequal}\ n} follows from the assumption, again using -the induction hypothesis:% -\end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast\ elim{\isacharcolon}\ less{\isacharunderscore}SucE{\isacharparenright}\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -325,7 +202,7 @@ \isamarkuptrue% \isacommand{theorem}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}insert\ induct{\isacharunderscore}lem{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% HOL already provides the mother of diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Misc/document/Itrev.tex --- a/doc-src/TutorialI/Misc/document/Itrev.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Misc/document/Itrev.tex Tue Feb 01 18:01:57 2005 +0100 @@ -65,58 +65,12 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -There is no choice as to the induction variable, and we immediately simplify:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Unfortunately, this attempt does not prove -the induction step: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }itrev\ list\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ rev\ list\ {\isasymLongrightarrow}\ itrev\ list\ {\isacharbrackleft}a{\isacharbrackright}\ {\isacharequal}\ rev\ list\ {\isacharat}\ {\isacharbrackleft}a{\isacharbrackright}% -\end{isabelle} -The induction hypothesis is too weak. The fixed -argument,~\isa{{\isacharbrackleft}{\isacharbrackright}}, prevents it from rewriting the conclusion. -This example suggests a heuristic: -\begin{quote}\index{generalizing induction formulae}% -\emph{Generalize goals for induction by replacing constants by variables.} -\end{quote} -Of course one cannot do this na\"{\i}vely: \isa{itrev\ xs\ ys\ {\isacharequal}\ rev\ xs} is -just not true. The correct generalization is% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -If \isa{ys} is replaced by \isa{{\isacharbrackleft}{\isacharbrackright}}, the right-hand side simplifies to -\isa{rev\ xs}, as required. - -In this instance it was easy to guess the right generalization. -Other situations can require a good deal of creativity. - -Although we now have two variables, only \isa{xs} is suitable for -induction, and we repeat our proof attempt. Unfortunately, we are still -not there: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }itrev\ list\ ys\ {\isacharequal}\ rev\ list\ {\isacharat}\ ys\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }itrev\ list\ {\isacharparenleft}a\ {\isacharhash}\ ys{\isacharparenright}\ {\isacharequal}\ rev\ list\ {\isacharat}\ a\ {\isacharhash}\ ys% -\end{isabelle} -The induction hypothesis is still too weak, but this time it takes no -intuition to generalize: the problem is that \isa{ys} is fixed throughout -the subgoal, but the induction hypothesis needs to be applied with -\isa{a\ {\isacharhash}\ ys} instead of \isa{ys}. Hence we prove the theorem -for all \isa{ys} instead of a fixed one:% -\end{isamarkuptxt}% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}ys{\isachardot}\ itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}\isamarkupfalse% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Misc/document/case_exprs.tex --- a/doc-src/TutorialI/Misc/document/case_exprs.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Misc/document/case_exprs.tex Tue Feb 01 18:01:57 2005 +0100 @@ -67,20 +67,9 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -results in the proof state -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ {\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ xs{\isacharparenright}\ {\isacharequal}\ xs\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }xs\ {\isacharequal}\ a\ {\isacharhash}\ list\ {\isasymLongrightarrow}\ {\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ xs{\isacharparenright}\ {\isacharequal}\ xs% -\end{isabelle} -which is solved automatically:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Misc/document/natsum.tex --- a/doc-src/TutorialI/Misc/document/natsum.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Misc/document/natsum.tex Tue Feb 01 18:01:57 2005 +0100 @@ -24,11 +24,9 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}sum\ n\ {\isacharplus}\ sum\ n\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \newcommand{\mystar}{*% @@ -109,7 +107,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}min\ i\ {\isacharparenleft}max\ j\ {\isacharparenleft}k{\isacharasterisk}k{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ max\ {\isacharparenleft}min\ {\isacharparenleft}k{\isacharasterisk}k{\isacharparenright}\ i{\isacharparenright}\ {\isacharparenleft}min\ i\ {\isacharparenleft}j{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}arith{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Misc/document/simp.tex --- a/doc-src/TutorialI/Misc/document/simp.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Misc/document/simp.tex Tue Feb 01 18:01:57 2005 +0100 @@ -114,9 +114,8 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isasymlbrakk}\ xs\ {\isacharat}\ zs\ {\isacharequal}\ ys\ {\isacharat}\ xs{\isacharsemicolon}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ ys\ {\isacharequal}\ zs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -128,19 +127,9 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -An unmodified application of \isa{simp} loops. The culprit is the -simplification rule \isa{f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}}, which is extracted from -the assumption. (Isabelle notices certain simple forms of -nontermination but not this one.) The problem can be circumvented by -telling the simplifier to ignore the assumptions:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -198,27 +187,12 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}xor\ A\ {\isacharparenleft}{\isasymnot}A{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Typically, we begin by unfolding some definitions: -\indexbold{definitions!unfolding}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -In this particular case, the resulting goal -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ A\ {\isasymand}\ {\isasymnot}\ {\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ A\ {\isasymand}\ {\isasymnot}\ A% -\end{isabelle} -can be proved by simplification. Thus we could have proved the lemma outright by% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -255,9 +229,8 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}let\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ in\ xs{\isacharat}ys{\isacharat}xs{\isacharparenright}\ {\isacharequal}\ ys{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% If, in a particular context, there is no danger of a combinatorial explosion @@ -279,9 +252,8 @@ \isamarkuptrue% \isacommand{lemma}\ hd{\isacharunderscore}Cons{\isacharunderscore}tl{\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ \ {\isasymLongrightarrow}\ \ hd\ xs\ {\isacharhash}\ tl\ xs\ {\isacharequal}\ xs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -317,51 +289,17 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}xs{\isachardot}\ if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ else\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The goal can be split by a special method, \methdx{split}:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymforall}xs{\isachardot}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}% -\end{isabelle} -where \tdx{split_if} is a theorem that expresses splitting of -\isa{if}s. Because -splitting the \isa{if}s is usually the right proof strategy, the -simplifier does it automatically. Try \isacommand{apply}\isa{{\isacharparenleft}simp{\isacharparenright}} -on the initial goal above. - -This splitting idea generalizes from \isa{if} to \sdx{case}. -Let us simplify a case analysis over lists:\index{*list.split (theorem)}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ zs\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ y{\isacharhash}{\isacharparenleft}ys{\isacharat}zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ xs{\isacharat}zs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ list{\isachardot}split{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isacharparenright}\ {\isasymand}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ }{\isacharparenleft}{\isasymforall}a\ list{\isachardot}\ xs\ {\isacharequal}\ a\ {\isacharhash}\ list\ {\isasymlongrightarrow}\ a\ {\isacharhash}\ list\ {\isacharat}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isacharparenright}% -\end{isabelle} -The simplifier does not split -\isa{case}-expressions, as it does \isa{if}-expressions, -because with recursive datatypes it could lead to nontermination. -Instead, the simplifier has a modifier -\isa{split}\index{*split (modifier)} -for adding splitting rules explicitly. The -lemma above can be proved in one step by% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ split{\isacharcolon}\ list{\isachardot}split{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -382,7 +320,7 @@ \end{isamarkuptext}% \isamarkuptrue% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ split\ del{\isacharcolon}\ split{\isacharunderscore}if{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -407,30 +345,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ else\ ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharunderscore}asm{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Unlike splitting the conclusion, this step creates two -separate subgoals, which here can be solved by \isa{simp{\isacharunderscore}all}: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}% -\end{isabelle} -If you need to split both in the assumptions and the conclusion, -use $t$\isa{{\isachardot}splits} which subsumes $t$\isa{{\isachardot}split} and -$t$\isa{{\isachardot}split{\isacharunderscore}asm}. Analogously, there is \isa{if{\isacharunderscore}splits}. - -\begin{warn} - The simplifier merely simplifies the condition of an - \isa{if}\index{*if expressions!simplification of} but not the - \isa{then} or \isa{else} parts. The latter are simplified only after the - condition reduces to \isa{True} or \isa{False}, or after splitting. The - same is true for \sdx{case}-expressions: only the selector is - simplified at first, until either the expression reduces to one of the - cases or it is split. -\end{warn}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% % @@ -446,11 +361,10 @@ on:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{ML}\ {\isachardoublequote}set\ trace{\isacharunderscore}simp{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}rev\ {\isacharbrackleft}a{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -487,7 +401,6 @@ rules. Thus it is advisable to reset it:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{ML}\ {\isachardoublequote}reset\ trace{\isacharunderscore}simp{\isachardoublequote}\isanewline \isamarkupfalse% \isamarkupfalse% \end{isabellebody}% diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Recdef/document/Induction.tex --- a/doc-src/TutorialI/Recdef/document/Induction.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Recdef/document/Induction.tex Tue Feb 01 18:01:57 2005 +0100 @@ -21,33 +21,11 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}map\ f\ {\isacharparenleft}sep{\isacharparenleft}x{\isacharcomma}xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep{\isacharparenleft}f\ x{\isacharcomma}\ map\ f\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Note that \isa{map\ f\ xs} -is the result of applying \isa{f} to all elements of \isa{xs}. We prove -this lemma by recursion induction over \isa{sep}:% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The resulting proof state has three subgoals corresponding to the three -clauses for \isa{sep}: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ x{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ x\ y\ zs{\isachardot}\isanewline -\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}% -\end{isabelle} -The rest is pure simplification:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ simp{\isacharunderscore}all\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Try proving the above lemma by structural induction, and you find that you diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Recdef/document/Nested2.tex --- a/doc-src/TutorialI/Recdef/document/Nested2.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Recdef/document/Nested2.tex Tue Feb 01 18:01:57 2005 +0100 @@ -5,7 +5,7 @@ \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ {\isacharless}\ Suc{\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -20,19 +20,9 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}trev{\isacharparenleft}trev\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ rule{\isacharcolon}\ trev{\isachardot}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ trev\ {\isacharparenleft}trev\ {\isacharparenleft}Var\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ Var\ x\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}f\ ts{\isachardot}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymforall}x{\isachardot}\ x\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ trev\ {\isacharparenleft}trev\ x{\isacharparenright}\ {\isacharequal}\ x\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }trev\ {\isacharparenleft}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ App\ f\ ts% -\end{isabelle} -Both the base case and the induction step fall to simplification:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{by}{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ rev{\isacharunderscore}map\ sym{\isacharbrackleft}OF\ map{\isacharunderscore}compose{\isacharbrackright}\ cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Recdef/document/simplification.tex --- a/doc-src/TutorialI/Recdef/document/simplification.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Recdef/document/simplification.tex Tue Feb 01 18:01:57 2005 +0100 @@ -94,16 +94,14 @@ \isamarkuptrue% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}gcd\ {\isacharparenleft}m{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ m{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ gcd{\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Recdef/document/termination.tex --- a/doc-src/TutorialI/Recdef/document/termination.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Recdef/document/termination.tex Tue Feb 01 18:01:57 2005 +0100 @@ -65,9 +65,8 @@ \isamarkuptrue% \isacommand{theorem}\ {\isachardoublequote}qs{\isacharbrackleft}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharbrackright}\ {\isacharequal}\ qs{\isacharbrackleft}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{2}}{\isacharbrackright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/ToyList/document/ToyList.tex --- a/doc-src/TutorialI/ToyList/document/ToyList.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/ToyList/document/ToyList.tex Tue Feb 01 18:01:57 2005 +0100 @@ -136,98 +136,10 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{theorem}\ rev{\isacharunderscore}rev\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\index{theorem@\isacommand {theorem} (command)|bold}% -\noindent -This \isacommand{theorem} command does several things: -\begin{itemize} -\item -It establishes a new theorem to be proved, namely \isa{rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs}. -\item -It gives that theorem the name \isa{rev{\isacharunderscore}rev}, for later reference. -\item -It tells Isabelle (via the bracketed attribute \attrdx{simp}) to take the eventual theorem as a simplification rule: future proofs involving -simplification will replace occurrences of \isa{rev\ {\isacharparenleft}rev\ xs{\isacharparenright}} by -\isa{xs}. -\end{itemize} -The name and the simplification attribute are optional. -Isabelle's response is to print the initial proof state consisting -of some header information (like how many subgoals there are) followed by -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs% -\end{isabelle} -For compactness reasons we omit the header in this tutorial. -Until we have finished a proof, the \rmindex{proof state} proper -always looks like this: -\begin{isabelle} -~1.~$G\sb{1}$\isanewline -~~\vdots~~\isanewline -~$n$.~$G\sb{n}$ -\end{isabelle} -The numbered lines contain the subgoals $G\sb{1}$, \dots, $G\sb{n}$ -that we need to prove to establish the main goal.\index{subgoals} -Initially there is only one subgoal, which is identical with the -main goal. (If you always want to see the main goal as well, -set the flag \isa{Proof.show_main_goal}\index{*show_main_goal (flag)} ---- this flag used to be set by default.) - -Let us now get back to \isa{rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs}. Properties of recursively -defined functions are best established by induction. In this case there is -nothing obvious except induction on \isa{xs}:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent\index{*induct_tac (method)}% -This tells Isabelle to perform induction on variable \isa{xs}. The suffix -\isa{tac} stands for \textbf{tactic},\index{tactics} -a synonym for ``theorem proving function''. -By default, induction acts on the first subgoal. The new proof state contains -two subgoals, namely the base case (\isa{Nil}) and the induction step -(\isa{Cons}): -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ rev\ {\isacharparenleft}rev\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }rev\ {\isacharparenleft}rev\ list{\isacharparenright}\ {\isacharequal}\ list\ {\isasymLongrightarrow}\ rev\ {\isacharparenleft}rev\ {\isacharparenleft}a\ {\isacharhash}\ list{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ a\ {\isacharhash}\ list% -\end{isabelle} - -The induction step is an example of the general format of a subgoal:\index{subgoals} -\begin{isabelle} -~$i$.~{\isasymAnd}$x\sb{1}$~\dots$x\sb{n}$.~{\it assumptions}~{\isasymLongrightarrow}~{\it conclusion} -\end{isabelle}\index{$IsaAnd@\isasymAnd|bold} -The prefix of bound variables \isasymAnd$x\sb{1}$~\dots~$x\sb{n}$ can be -ignored most of the time, or simply treated as a list of variables local to -this subgoal. Their deeper significance is explained in Chapter~\ref{chap:rules}. -The {\it assumptions}\index{assumptions!of subgoal} -are the local assumptions for this subgoal and {\it - conclusion}\index{conclusion!of subgoal} is the actual proposition to be proved. -Typical proof steps -that add new assumptions are induction and case distinction. In our example -the only assumption is the induction hypothesis \isa{rev\ {\isacharparenleft}rev\ list{\isacharparenright}\ {\isacharequal}\ list}, where \isa{list} is a variable name chosen by Isabelle. If there -are multiple assumptions, they are enclosed in the bracket pair -\indexboldpos{\isasymlbrakk}{$Isabrl} and -\indexboldpos{\isasymrbrakk}{$Isabrr} and separated by semicolons. - -Let us try to solve both goals automatically:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This command tells Isabelle to apply a proof strategy called -\isa{auto} to all subgoals. Essentially, \isa{auto} tries to -simplify the subgoals. In our case, subgoal~1 is solved completely (thanks -to the equation \isa{rev\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}) and disappears; the simplified version -of subgoal~2 becomes the new subgoal~1: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }rev\ {\isacharparenleft}rev\ list{\isacharparenright}\ {\isacharequal}\ list\ {\isasymLongrightarrow}\ rev\ {\isacharparenleft}rev\ list\ {\isacharat}\ a\ {\isacharhash}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ a\ {\isacharhash}\ list% -\end{isabelle} -In order to simplify this subgoal further, a lemma suggests itself.% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% % @@ -242,35 +154,10 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent The keywords \commdx{theorem} and -\commdx{lemma} are interchangeable and merely indicate -the importance we attach to a proposition. Therefore we use the words -\emph{theorem} and \emph{lemma} pretty much interchangeably, too. - -There are two variables that we could induct on: \isa{xs} and -\isa{ys}. Because \isa{{\isacharat}} is defined by recursion on -the first argument, \isa{xs} is the correct one:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This time not even the base case is solved automatically:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ rev\ ys\ {\isacharequal}\ rev\ ys\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}% -\end{isabelle} -Again, we need to abandon this proof attempt and prove another simple lemma -first. In the future the step of abandoning an incomplete proof before -embarking on the proof of a lemma usually remains implicit.% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% % @@ -284,21 +171,10 @@ \isamarkuptrue% \isacommand{lemma}\ app{\isacharunderscore}Nil{\isadigit{2}}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -It works, yielding the desired message \isa{No\ subgoals{\isacharbang}}: -\begin{isabelle}% -xs\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs\isanewline -No\ subgoals{\isacharbang}% -\end{isabelle} -We still need to confirm that the proof is now finished:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -318,27 +194,8 @@ \isamarkuptrue% \isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -we find that this time \isa{auto} solves the base case, but the -induction step merely simplifies to -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ list{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }rev\ {\isacharparenleft}list\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ rev\ ys\ {\isacharat}\ rev\ list\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isacharparenleft}rev\ ys\ {\isacharat}\ rev\ list{\isacharparenright}\ {\isacharat}\ a\ {\isacharhash}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ rev\ ys\ {\isacharat}\ rev\ list\ {\isacharat}\ a\ {\isacharhash}\ {\isacharbrackleft}{\isacharbrackright}% -\end{isabelle} -Now we need to remember that \isa{{\isacharat}} associates to the right, and that -\isa{{\isacharhash}} and \isa{{\isacharat}} have the same priority (namely the \isa{{\isadigit{6}}{\isadigit{5}}} -in their \isacommand{infixr} annotation). Thus the conclusion really is -\begin{isabelle} -~~~~~(rev~ys~@~rev~list)~@~(a~\#~[])~=~rev~ys~@~(rev~list~@~(a~\#~[])) -\end{isabelle} -and the missing lemma is associativity of \isa{{\isacharat}}.% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% % @@ -353,11 +210,9 @@ \isamarkuptrue% \isacommand{lemma}\ app{\isacharunderscore}assoc\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharat}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ {\isacharparenleft}ys\ {\isacharat}\ zs{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -366,11 +221,9 @@ \isamarkuptrue% \isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -379,11 +232,9 @@ \isamarkuptrue% \isacommand{theorem}\ rev{\isacharunderscore}rev\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Trie/document/Trie.tex --- a/doc-src/TutorialI/Trie/document/Trie.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Trie/document/Trie.tex Tue Feb 01 18:01:57 2005 +0100 @@ -63,9 +63,8 @@ \isamarkuptrue% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}lookup\ {\isacharparenleft}Trie\ None\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ as\ {\isacharequal}\ None{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ as{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Things begin to get interesting with the definition of an update function @@ -110,36 +109,11 @@ \isamarkuptrue% \isacommand{theorem}\ {\isachardoublequote}{\isasymforall}t\ v\ bs{\isachardot}\ lookup\ {\isacharparenleft}update\ t\ as\ v{\isacharparenright}\ bs\ {\isacharequal}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}if\ as{\isacharequal}bs\ then\ Some\ v\ else\ lookup\ t\ bs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Our plan is to induct on \isa{as}; hence the remaining variables are -quantified. From the definitions it is clear that induction on either -\isa{as} or \isa{bs} is required. The choice of \isa{as} is -guided by the intuition that simplification of \isa{lookup} might be easier -if \isa{update} has already been simplified, which can only happen if -\isa{as} is instantiated. -The start of the proof is conventional:% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ as{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Unfortunately, this time we are left with three intimidating looking subgoals: -\begin{isabelle} -~1.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs\isanewline -~2.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs\isanewline -~3.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs -\end{isabelle} -Clearly, if we want to make headway we have to instantiate \isa{bs} as -well now. It turns out that instead of induction, case distinction -suffices:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac{\isacharbrackleft}{\isacharbang}{\isacharbrackright}\ bs{\isacharcomma}\ auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Axioms.tex --- a/doc-src/TutorialI/Types/document/Axioms.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Axioms.tex Tue Feb 01 18:01:57 2005 +0100 @@ -49,22 +49,8 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}parord{\isacharparenright}\ {\isacharless}{\isacharless}\ y\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymnot}\ y\ {\isacharless}{\isacharless}\ x{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -The conclusion is not just \isa{{\isasymnot}\ y\ {\isacharless}{\isacharless}\ x} because the -simplifier's preprocessor (see \S\ref{sec:simp-preprocessor}) -would turn it into \isa{{\isacharparenleft}y\ {\isacharless}{\isacharless}\ x{\isacharparenright}\ {\isacharequal}\ False}, yielding -a nonterminating rewrite rule. -(It would be used to try to prove its own precondition \emph{ad - infinitum}.) -In the form above, the rule is useful. -The type constraint is necessary because otherwise Isabelle would only assume -\isa{{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel} (as required in the type of \isa{{\isacharless}{\isacharless}}), -when the proposition is not a theorem. The proof is easy:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ less{\isacharunderscore}le\ antisym{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% We could now continue in this vein and develop a whole theory of @@ -76,26 +62,10 @@ \isamarkuptrue% \isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ parord\isanewline \isamarkupfalse% -\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -This time \isa{intro{\isacharunderscore}classes} leaves us with the four axioms, -specialized to type \isa{bool}, as subgoals: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isasymColon}bool{\isachardot}\ x\ {\isacharless}{\isacharless}{\isacharequal}\ x\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}{\isacharparenleft}x{\isasymColon}bool{\isacharparenright}\ {\isacharparenleft}y{\isasymColon}bool{\isacharparenright}\ z{\isasymColon}bool{\isachardot}\ {\isasymlbrakk}x\ {\isacharless}{\isacharless}{\isacharequal}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharless}{\isacharless}{\isacharequal}\ z\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}{\isacharparenleft}x{\isasymColon}bool{\isacharparenright}\ y{\isasymColon}bool{\isachardot}\ {\isasymlbrakk}x\ {\isacharless}{\isacharless}{\isacharequal}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ y\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}{\isacharparenleft}x{\isasymColon}bool{\isacharparenright}\ y{\isasymColon}bool{\isachardot}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}% -\end{isabelle} -Fortunately, the proof is easy for \isa{blast} -once we have unfolded the definitions -of \isa{{\isacharless}{\isacharless}} and \isa{{\isacharless}{\isacharless}{\isacharequal}} at type \isa{bool}:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}\ only{\isacharcolon}\ le{\isacharunderscore}bool{\isacharunderscore}def\ less{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -106,7 +76,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}P{\isacharcolon}{\isacharcolon}bool{\isacharparenright}\ {\isacharless}{\isacharless}\ Q\ {\isasymLongrightarrow}\ {\isasymnot}{\isacharparenleft}Q\ {\isacharless}{\isacharless}\ P{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -138,13 +108,10 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isasymAnd}x{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}linord{\isachardot}\ x\ {\isacharless}{\isacharless}\ y\ {\isasymor}\ x\ {\isacharequal}\ y\ {\isasymor}\ y\ {\isacharless}{\isacharless}\ x{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ less{\isacharunderscore}le{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}insert\ linear{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Linear orders are an example of subclassing\index{subclasses} @@ -178,28 +145,12 @@ \isamarkuptrue% \isacommand{instance}\ parord\ {\isacharless}\ strord\isanewline \isamarkupfalse% -\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isasymColon}{\isacharprime}a{\isachardot}\ {\isasymnot}\ x\ {\isacharless}{\isacharless}\ x\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}{\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharparenleft}y{\isasymColon}{\isacharprime}a{\isacharparenright}\ z{\isasymColon}{\isacharprime}a{\isachardot}\ {\isasymlbrakk}x\ {\isacharless}{\isacharless}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharless}{\isacharless}\ z\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}{\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ y{\isasymColon}{\isacharprime}a{\isachardot}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}{\isacharequal}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}\ y\ {\isasymor}\ x\ {\isacharequal}\ y{\isacharparenright}\isanewline -type\ variables{\isacharcolon}\isanewline -\isaindent{\ \ }{\isacharprime}a\ {\isacharcolon}{\isacharcolon}\ parord% -\end{isabelle} -Assuming \isa{{\isacharprime}a\ {\isacharcolon}{\isacharcolon}\ parord}, the three axioms of class \isa{strord} -are easily proved:% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}\ add{\isacharcolon}\ less{\isacharunderscore}le{\isacharparenright}\isanewline -\ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ trans\ antisym{\isacharparenright}\isanewline +\isamarkupfalse% +\isamarkuptrue% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ refl{\isacharparenright}\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The subclass relation must always be acyclic. Therefore Isabelle will diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Numbers.tex --- a/doc-src/TutorialI/Types/document/Numbers.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Numbers.tex Tue Feb 01 18:01:57 2005 +0100 @@ -5,23 +5,16 @@ \isacommand{theory}\ Numbers\ {\isacharequal}\ Real{\isacharcolon}\isanewline \isanewline \isamarkupfalse% -\isacommand{ML}\ {\isachardoublequote}Pretty{\isachardot}setmargin\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{ML}\ {\isachardoublequote}IsarOutput{\isachardot}indent\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% numeric literals; default simprules; can re-orient% \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ m\ {\isacharequal}\ m\ {\isacharplus}\ m{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isacharparenleft}{\isadigit{2}}{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharasterisk}\ m\ {\isacharequal}\ m\ {\isacharplus}\ m% -\end{isabelle}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{oops}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{consts}\ h\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline @@ -75,23 +68,10 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}Suc{\isacharparenleft}i\ {\isacharplus}\ j{\isacharasterisk}l{\isacharasterisk}k\ {\isacharplus}\ m{\isacharasterisk}n{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}n{\isacharasterisk}m\ {\isacharplus}\ i\ {\isacharplus}\ k{\isacharasterisk}j{\isacharasterisk}l{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ Suc\ {\isacharparenleft}i\ {\isacharplus}\ j\ {\isacharasterisk}\ l\ {\isacharasterisk}\ k\ {\isacharplus}\ m\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}n\ {\isacharasterisk}\ m\ {\isacharplus}\ i\ {\isacharplus}\ k\ {\isacharasterisk}\ j\ {\isacharasterisk}\ l{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ add{\isacharunderscore}ac\ mult{\isacharunderscore}ac{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ Suc\ {\isacharparenleft}i\ {\isacharplus}\ {\isacharparenleft}m\ {\isacharasterisk}\ n\ {\isacharplus}\ j\ {\isacharasterisk}\ {\isacharparenleft}k\ {\isacharasterisk}\ l{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ }f\ {\isacharparenleft}i\ {\isacharplus}\ {\isacharparenleft}m\ {\isacharasterisk}\ n\ {\isacharplus}\ j\ {\isacharasterisk}\ {\isacharparenleft}k\ {\isacharasterisk}\ l{\isacharparenright}{\isacharparenright}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{oops}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -117,33 +97,17 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ n\ {\isacharasterisk}\ n\ {\isacharminus}\ {\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}clarsimp\ split{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split\ iff\ del{\isacharcolon}\ less{\isacharunderscore}Suc{\isadigit{0}}{\isacharparenright}\isanewline -\ % -\isamarkupcmt{\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}d{\isachardot}\ {\isasymlbrakk}n\ {\isacharless}\ Suc\ {\isadigit{0}}{\isacharsemicolon}\ n\ {\isacharasterisk}\ n\ {\isacharequal}\ Suc\ d{\isasymrbrakk}\ {\isasymLongrightarrow}\ d\ {\isacharequal}\ {\isadigit{0}}% -\end{isabelle}% -} +\isamarkupfalse% +\isamarkupfalse% \isanewline -\isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequote}n{\isacharequal}{\isadigit{0}}{\isachardoublequote}{\isacharcomma}\ force{\isacharcomma}\ arith{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{done}\isanewline \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}n\ {\isacharminus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ n\ {\isacharasterisk}\ n\ {\isacharminus}\ {\isacharparenleft}{\isadigit{4}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}simp\ split{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharcomma}\ clarify{\isacharparenright}\isanewline -\ % -\isamarkupcmt{\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}d{\isachardot}\ {\isasymlbrakk}n\ {\isacharless}\ {\isadigit{2}}{\isacharsemicolon}\ n\ {\isacharasterisk}\ n\ {\isacharequal}\ {\isadigit{4}}\ {\isacharplus}\ d{\isasymrbrakk}\ {\isasymLongrightarrow}\ d\ {\isacharequal}\ {\isadigit{0}}% -\end{isabelle}% -} -\isanewline +\isamarkupfalse% \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequote}n{\isacharequal}{\isadigit{0}}\ {\isacharbar}\ n{\isacharequal}{\isadigit{1}}{\isachardoublequote}{\isacharcomma}\ force{\isacharcomma}\ arith{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -264,12 +228,12 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}abs\ {\isacharparenleft}x{\isacharplus}y{\isacharparenright}\ {\isasymle}\ abs\ x\ {\isacharplus}\ abs\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ int{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ arith\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}abs\ {\isacharparenleft}{\isadigit{2}}{\isacharasterisk}x{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ abs\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ int{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ zabs{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Induction rules for the Integers @@ -344,52 +308,26 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{3}}{\isacharslash}{\isadigit{4}}\ {\isacharless}\ {\isacharparenleft}{\isadigit{7}}{\isacharslash}{\isadigit{8}}\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\ \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}P\ {\isacharparenleft}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ P\ {\isacharparenleft}{\isadigit{3}}\ {\isacharslash}\ {\isadigit{4}}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}\ {\isacharslash}\ {\isadigit{1}}{\isadigit{5}}{\isacharparenright}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}\ simp\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ P\ {\isacharparenleft}{\isadigit{2}}\ {\isacharslash}\ {\isadigit{5}}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{oops}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{3}}\ {\isacharslash}\ {\isadigit{4}}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}\ {\isacharslash}\ {\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ x% -\end{isabelle}% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}\ simp\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isadigit{2}}\ {\isacharless}\ x\ {\isacharasterisk}\ {\isadigit{5}}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{oops}\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{0}}{\isacharcircum}{\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ simp\ \isanewline \isamarkupfalse% -\isacommand{oops}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Ring and Field @@ -417,7 +355,7 @@ \rulename{field_mult_cancel_right}% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}set\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% effect of show sorts on the above @@ -429,7 +367,7 @@ \rulename{field_mult_cancel_right}% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}reset\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% absolute value diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Overloading.tex --- a/doc-src/TutorialI/Types/document/Overloading.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Overloading.tex Tue Feb 01 18:01:57 2005 +0100 @@ -4,7 +4,7 @@ \isamarkupfalse% \isacommand{instance}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}ordrel\isanewline \isamarkupfalse% -\isacommand{by}\ intro{\isacharunderscore}classes\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Overloading1.tex --- a/doc-src/TutorialI/Types/document/Overloading1.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Overloading1.tex Tue Feb 01 18:01:57 2005 +0100 @@ -41,16 +41,8 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ ordrel\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -Command \isacommand{instance} actually starts a proof, namely that -\isa{bool} satisfies all axioms of \isa{ordrel}. -There are none, but we still need to finish that proof, which we do -by invoking the \methdx{intro_classes} method:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{by}\ intro{\isacharunderscore}classes\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -72,7 +64,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}False\ {\isacharless}{\isacharless}{\isacharequal}\ P{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Overloading2.tex --- a/doc-src/TutorialI/Types/document/Overloading2.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Overloading2.tex Tue Feb 01 18:01:57 2005 +0100 @@ -12,7 +12,7 @@ \isamarkuptrue% \isacommand{instance}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ordrel{\isacharparenright}ordrel\isanewline \isamarkupfalse% -\isacommand{by}\ intro{\isacharunderscore}classes\isanewline +\isanewline \isanewline \isamarkupfalse% \isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Pairs.tex --- a/doc-src/TutorialI/Types/document/Pairs.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Pairs.tex Tue Feb 01 18:01:57 2005 +0100 @@ -67,7 +67,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}x{\isacharparenright}\ p\ {\isacharequal}\ fst\ p{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% This works well if rewriting with \isa{split{\isacharunderscore}def} finishes the @@ -90,19 +90,11 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}y{\isacharparenright}\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymforall}x\ y{\isachardot}\ p\ {\isacharequal}\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymlongrightarrow}\ y\ {\isacharequal}\ snd\ p% -\end{isabelle} -This subgoal is easily proved by simplification. Thus we could have combined -simplification and splitting in one command that proves the goal outright:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ split{\isacharcolon}\ split{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Let us look at a second example:% @@ -110,42 +102,15 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}let\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isacharequal}\ p\ in\ fst\ p\ {\isacharequal}\ x{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\ fst\ p\ {\isacharequal}\ x{\isacharparenright}\ p% -\end{isabelle} -A paired \isa{let} reduces to a paired $\lambda$-abstraction, which -can be split as above. The same is true for paired set comprehension:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}p\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y{\isacharbraceright}\ {\isasymlongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}\ simp\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ split\ op\ {\isacharequal}\ p\ {\isasymlongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p% -\end{isabelle} -Again, simplification produces a term suitable for \isa{split{\isacharunderscore}split} -as above. If you are worried about the strange form of the premise: -\isa{{\isasymlambda}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\ x\ {\isacharequal}\ y} is short for \isa{{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y}. -The same proof procedure works for% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}p\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y{\isacharbraceright}\ {\isasymLongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -except that we now have to use \isa{split{\isacharunderscore}split{\isacharunderscore}asm}, because -\isa{split} occurs in the assumptions. - -However, splitting \isa{split} is not always a solution, as no \isa{split} -may be present in the goal. Consider the following function:% -\end{isamarkuptxt}% \isamarkuptrue% \isamarkupfalse% \isacommand{consts}\ swap\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymtimes}\ {\isacharprime}a{\isachardoublequote}\isanewline @@ -160,50 +125,16 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}swap{\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ p{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -simplification will do nothing, because the defining equation for \isa{swap} -expects a pair. Again, we need to turn \isa{p} into a pair first, but this -time there is no \isa{split} in sight. In this case the only thing we can do -is to split the term by hand:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ p{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ b{\isachardot}\ p\ {\isacharequal}\ {\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isasymLongrightarrow}\ swap\ {\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ p% -\end{isabelle} -Again, \methdx{case_tac} is applicable because \isa{{\isasymtimes}} is a datatype. -The subgoal is easily proved by \isa{simp}. - -Splitting by \isa{case{\isacharunderscore}tac} also solves the previous examples and may thus -appear preferable to the more arcane methods introduced first. However, see -the warning about \isa{case{\isacharunderscore}tac} in \S\ref{sec:struct-ind-case}. - -In case the term to be split is a quantified variable, there are more options. -You can split \emph{all} \isa{{\isasymAnd}}-quantified variables in a goal -with the rewrite rule \isa{split{\isacharunderscore}paired{\isacharunderscore}all}:% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymAnd}p\ q{\isachardot}\ swap{\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ q\ {\isasymlongrightarrow}\ p\ {\isacharequal}\ q{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ b\ aa\ ba{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }swap\ {\isacharparenleft}swap\ {\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}aa{\isacharcomma}\ ba{\isacharparenright}\ {\isasymlongrightarrow}\ {\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}aa{\isacharcomma}\ ba{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -218,7 +149,7 @@ \end{isamarkuptext}% \isamarkuptrue% \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% \isamarkupfalse% % \begin{isamarkuptext}% @@ -229,7 +160,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}p{\isachardot}\ {\isasymexists}q{\isachardot}\ swap\ p\ {\isacharequal}\ swap\ q{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Records.tex --- a/doc-src/TutorialI/Types/document/Records.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Records.tex Tue Feb 01 18:01:57 2005 +0100 @@ -77,7 +77,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}Xcoord\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ a{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The \emph{update}\index{update!record} operation is functional. For @@ -88,7 +88,7 @@ \isacommand{lemma}\ {\isachardoublequote}{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isasymrparr}\ {\isacharequal}\isanewline \ \ \ \ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isadigit{0}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{warn} @@ -139,7 +139,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}Xcoord\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ p{\isasymrparr}\ {\isacharequal}\ a{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% This lemma applies to any record whose first two fields are \isa{Xcoord} and~\isa{Ycoord}. Note that \isa{{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenright}{\isasymrparr}} is exactly the same as \isa{{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}}. Selectors and updates are always polymorphic wrt.\ the @@ -152,7 +152,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}point{\isachardot}more\ cpt{\isadigit{1}}\ {\isacharequal}\ {\isasymlparr}col\ {\isacharequal}\ Green{\isasymrparr}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% We see that the colour part attached to this \isa{point} is a @@ -204,7 +204,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}incX\ r\ {\isacharequal}\ setX\ r\ {\isacharparenleft}getX\ r\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ getX{\isacharunderscore}def\ setX{\isacharunderscore}def\ incX{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{warn} @@ -231,7 +231,7 @@ \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharprime}{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharprime}{\isasymrparr}{\isacharparenright}\ {\isacharequal}\isanewline \ \ \ \ {\isacharparenleft}a\ {\isacharequal}\ a{\isacharprime}\ {\isasymand}\ b\ {\isacharequal}\ b{\isacharprime}{\isacharparenright}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The following equality is similar, but generic, in that \isa{r} @@ -240,7 +240,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharcolon}{\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Ycoord\ {\isacharcolon}{\isacharequal}\ b{\isacharcomma}\ Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% We see above the syntax for iterated updates. We could equivalently @@ -253,7 +253,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isasymrparr}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The generic version of this equality includes the pseudo-field @@ -262,7 +262,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ point{\isachardot}more\ r{\isasymrparr}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \medskip The simplifier can prove many record equalities @@ -272,9 +272,8 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{apply}\ simp{\isacharquery}\isanewline -\ \ \isamarkupfalse% -\isacommand{oops}\isamarkupfalse% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% Here the simplifier can do nothing, since general record equality is @@ -285,19 +284,10 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}drule{\isacharunderscore}tac\ f\ {\isacharequal}\ Xcoord\ \isakeyword{in}\ arg{\isacharunderscore}cong{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ Xcoord\ {\isacharparenleft}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}{\isacharparenright}\ {\isacharequal}\ Xcoord\ {\isacharparenleft}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}{\isacharparenright}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}% -\end{isabelle} - Now, \isa{simp} will reduce the assumption to the desired - conclusion.% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}\ simp\isanewline -\ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The \isa{cases} method is preferable to such a forward proof. We @@ -305,31 +295,11 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -The \methdx{cases} method adds an equality to replace the - named record term by an explicit record expression, listing all - fields. It even includes the pseudo-field \isa{more}, since the - record equality stated here is generic for all extensions.% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}cases\ r{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}Xcoord\ Ycoord\ more{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }r\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ more{\isasymrparr}{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}% -\end{isabelle} Again, \isa{simp} finishes the proof. Because \isa{r} is now represented as - an explicit record construction, the updates can be applied and the - record equality can be replaced by equality of the corresponding - fields (due to injectivity).% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}\ simp\isanewline -\ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% The generic cases method does not admit references to locally bound @@ -416,17 +386,10 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}cpt{\isadigit{1}}\ {\isacharequal}\ cpt{\isadigit{2}}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs\ cpoint{\isachardot}defs{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ Xcoord\ pt{\isadigit{1}}\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}\ {\isasymand}\ Ycoord\ pt{\isadigit{1}}\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}% -\end{isabelle}% -\end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ pt{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% +\isamarkuptrue% +\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% In the example below, a coloured point is truncated to leave a @@ -435,7 +398,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}point{\isachardot}truncate\ cpt{\isadigit{2}}\ {\isacharequal}\ pt{\isadigit{1}}{\isachardoublequote}\isanewline \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ pt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \begin{exercise} diff -r cb3612cc41a3 -r fc075ae929e4 doc-src/TutorialI/Types/document/Typedefs.tex --- a/doc-src/TutorialI/Types/document/Typedefs.tex Sun Jan 30 20:48:50 2005 +0100 +++ b/doc-src/TutorialI/Types/document/Typedefs.tex Tue Feb 01 18:01:57 2005 +0100 @@ -87,21 +87,9 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{typedef}\ three\ {\isacharequal}\ {\isachardoublequote}{\isacharbraceleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isacharcomma}\ {\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent -In order to enforce that the representing set on the right-hand side is -non-empty, this definition actually starts a proof to that effect: -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymexists}x{\isachardot}\ x\ {\isasymin}\ {\isacharbraceleft}{\isadigit{0}}{\isacharcomma}\ {\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharbraceright}% -\end{isabelle} -Fortunately, this is easy enough to show, even \isa{auto} could do it. -In general, one has to provide a witness, in our case 0:% -\end{isamarkuptxt}% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isadigit{0}}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% This type definition introduces the new type \isa{three} and asserts @@ -200,7 +188,7 @@ \isamarkuptrue% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymnoteq}\ B\ {\isasymand}\ B\ {\isasymnoteq}\ A\ {\isasymand}\ A\ {\isasymnoteq}\ C\ {\isasymand}\ C\ {\isasymnoteq}\ A\ {\isasymand}\ B\ {\isasymnoteq}\ C\ {\isasymand}\ C\ {\isasymnoteq}\ B{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}inject\ A{\isacharunderscore}def\ B{\isacharunderscore}def\ C{\isacharunderscore}def\ three{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -213,24 +201,11 @@ \end{isamarkuptext}% \isamarkuptrue% \isacommand{lemma}\ three{\isacharunderscore}cases{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ P\ A{\isacharsemicolon}\ P\ B{\isacharsemicolon}\ P\ C\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x{\isachardoublequote}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\noindent Again this follows easily from a pre-proved general theorem:% -\end{isamarkuptxt}% +\isamarkuptrue% +\isamarkupfalse% \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ rule{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}y{\isachardot}\ {\isasymlbrakk}P\ A{\isacharsemicolon}\ P\ B{\isacharsemicolon}\ P\ C{\isacharsemicolon}\ y\ {\isasymin}\ three{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Abs{\isacharunderscore}three\ y{\isacharparenright}% -\end{isabelle} -Simplification with \isa{three{\isacharunderscore}def} leads to the disjunction \isa{y\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ y\ {\isacharequal}\ {\isadigit{1}}\ {\isasymor}\ y\ {\isacharequal}\ {\isadigit{2}}} which \isa{auto} separates into three -subgoals, each of which is easily solved by simplification:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ three{\isacharunderscore}def\ A{\isacharunderscore}def\ B{\isacharunderscore}def\ C{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isamarkupfalse% % \begin{isamarkuptext}% \noindent diff -r cb3612cc41a3 -r fc075ae929e4 src/FOL/FOL.thy --- a/src/FOL/FOL.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/FOL/FOL.thy Tue Feb 01 18:01:57 2005 +0100 @@ -5,9 +5,12 @@ header {* Classical first-order logic *} -theory FOL = IFOL -files - ("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML"): +theory FOL +imports IFOL +files ("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML") + ("eqrule_FOL_data.ML") + ("~~/src/Provers/eqsubst.ML") +begin subsection {* The classical axiom *} @@ -43,6 +46,15 @@ setup Splitter.setup setup Clasimp.setup + +subsection {* Lucas Dixon's eqstep tactic *} + +use "~~/src/Provers/eqsubst.ML"; +use "eqrule_FOL_data.ML"; + +setup EQSubstTac.setup + + subsection {* Other simple lemmas *} lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)" diff -r cb3612cc41a3 -r fc075ae929e4 src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/FOL/IFOL.thy Tue Feb 01 18:01:57 2005 +0100 @@ -5,8 +5,10 @@ header {* Intuitionistic first-order logic *} -theory IFOL = Pure -files ("IFOL_lemmas.ML") ("fologic.ML") ("hypsubstdata.ML") ("intprover.ML"): +theory IFOL +imports Pure +files ("IFOL_lemmas.ML") ("fologic.ML") ("hypsubstdata.ML") ("intprover.ML") +begin subsection {* Syntax and axiomatic basis *} @@ -268,8 +270,6 @@ mp trans - - subsection {* ``Let'' declarations *} nonterminals letbinds letbind diff -r cb3612cc41a3 -r fc075ae929e4 src/FOL/IsaMakefile --- a/src/FOL/IsaMakefile Sun Jan 30 20:48:50 2005 +0100 +++ b/src/FOL/IsaMakefile Tue Feb 01 18:01:57 2005 +0100 @@ -32,6 +32,8 @@ $(SRC)/Provers/clasimp.ML $(SRC)/Provers/classical.ML \ $(SRC)/Provers/hypsubst.ML $(SRC)/Provers/ind.ML $(SRC)/Provers/induct_method.ML \ $(SRC)/Provers/simplifier.ML $(SRC)/Provers/splitter.ML $(SRC)/Provers/quantifier1.ML\ + $(SRC)/Provers/eqsubst.ML\ + eqrule_FOL_data.ML\ FOL.ML FOL.thy FOL_lemmas1.ML IFOL.ML IFOL.thy \ IFOL_lemmas.ML ROOT.ML blastdata.ML cladata.ML document/root.tex \ fologic.ML hypsubstdata.ML intprover.ML simpdata.ML diff -r cb3612cc41a3 -r fc075ae929e4 src/FOL/eqrule_FOL_data.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/eqrule_FOL_data.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,57 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/eqrule_FOL_data.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Created: 18 Feb 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Data for equality rules in the logic + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +structure ZF_EqRuleData : EQRULE_DATA = +struct + +fun mk_eq th = case concl_of th of + Const("==",_)$_$_ => Some (th) + | _$(Const("op <->",_)$_$_) => Some (th RS iff_reflection) + | _$(Const("op =",_)$_$_) => Some (th RS eq_reflection) + | _ => None; + +val tranformation_pairs = + [("op -->", [mp]), ("op &", [conjunct1,conjunct2]), + ("All", [spec]), ("True", []), ("False", [])]; + +(* +val mk_atomize: (string * thm list) list -> thm -> thm list +looks too specific to move it somewhere else +*) +fun mk_atomize pairs = + let fun atoms th = + (case Thm.concl_of th of + Const("Trueprop",_) $ p => + (case Term.head_of p of + Const(a,_) => + (case Library.assoc(pairs,a) of + Some(rls) => flat (map atoms ([th] RL rls)) + | None => [th]) + | _ => [th]) + | _ => [th]) + in atoms end; + +val prep_meta_eq = + (mapfilter + mk_eq + o (mk_atomize tranformation_pairs) + o Drule.gen_all + o zero_var_indexes) + +end; +structure EqRuleData = ZF_EqRuleData; + +structure EQSubstTac = + EQSubstTacFUN(structure EqRuleData = EqRuleData); diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Algebra/UnivPoly.thy Tue Feb 01 18:01:57 2005 +0100 @@ -357,7 +357,7 @@ case 0 then show ?case by (simp add: Pi_def) next case (Suc k) then show ?case - by (subst finsum_Suc2) (simp add: Pi_def a_comm)+ + by (simplesubst finsum_Suc2) (simp add: Pi_def a_comm)+ qed } note l = this diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Algebra/poly/LongDiv.thy --- a/src/HOL/Algebra/poly/LongDiv.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Algebra/poly/LongDiv.thy Tue Feb 01 18:01:57 2005 +0100 @@ -18,12 +18,10 @@ "!! f::(nat=>'a::ring). (ALL i. i < m --> f i = 0) --> setsum (%i. f (i+m)) {..d} = setsum f {..m+d}" apply (induct_tac d) - apply (induct_tac m) - apply (simp) - apply (force) - apply (simp) - apply (subst ab_semigroup_add.add_commute[of m]) - apply (simp) + apply (induct_tac m) + apply (simp) + apply (force) + apply (simp add: ab_semigroup_add.add_commute[of m]) done end diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Algebra/poly/UnivPoly2.thy --- a/src/HOL/Algebra/poly/UnivPoly2.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Algebra/poly/UnivPoly2.thy Tue Feb 01 18:01:57 2005 +0100 @@ -7,7 +7,9 @@ header {* Univariate Polynomials *} -theory UnivPoly2 = Abstract: +theory UnivPoly2 +imports "../abstract/Abstract" +begin (* already proved in Finite_Set.thy diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Complex/Complex.thy --- a/src/HOL/Complex/Complex.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Complex/Complex.thy Tue Feb 01 18:01:57 2005 +0100 @@ -865,9 +865,7 @@ iszero_def) lemma complex_number_of_cnj [simp]: "cnj(number_of v :: complex) = number_of v" -apply (subst complex_number_of [symmetric]) -apply (rule complex_cnj_complex_of_real) -done +by (simp only: complex_number_of [symmetric] complex_cnj_complex_of_real) lemma complex_number_of_cmod: "cmod(number_of v :: complex) = abs (number_of v :: real)" diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/HOL.thy Tue Feb 01 18:01:57 2005 +0100 @@ -8,6 +8,8 @@ theory HOL imports CPure files ("cladata.ML") ("blastdata.ML") ("simpdata.ML") ("antisym_setup.ML") + ("eqrule_HOL_data.ML") + ("~~/src/Provers/eqsubst.ML") begin subsection {* Primitive logic *} @@ -898,7 +900,7 @@ setup Blast.setup -subsubsection {* Simplifier setup *} +subsection {* Simplifier setup *} lemma meta_eq_to_obj_eq: "x == y ==> x = y" proof - @@ -1068,12 +1070,12 @@ lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" apply (rule case_split [of Q]) - apply (subst if_P) - prefer 3 apply (subst if_not_P, blast+) + apply (simplesubst if_P) + prefer 3 apply (simplesubst if_not_P, blast+) done lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" -by (subst split_if, blast) +by (simplesubst split_if, blast) lemmas if_splits = split_if split_if_asm @@ -1081,10 +1083,10 @@ by (rule split_if) lemma if_cancel: "(if c then x else x) = x" -by (subst split_if, blast) +by (simplesubst split_if, blast) lemma if_eq_cancel: "(if x = y then y else x) = x" -by (subst split_if, blast) +by (simplesubst split_if, blast) lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" -- {* This form is useful for expanding @{text if}s on the RIGHT of the @{text "==>"} symbol. *} @@ -1092,7 +1094,7 @@ lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" -- {* And this form is useful for expanding @{text if}s on the LEFT. *} - apply (subst split_if, blast) + apply (simplesubst split_if, blast) done lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) rules @@ -1113,11 +1115,23 @@ setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup setup Splitter.setup setup Clasimp.setup + +subsubsection {* Lucas Dixon's eqstep tactic *} + +use "~~/src/Provers/eqsubst.ML"; +use "eqrule_HOL_data.ML"; + +setup EQSubstTac.setup + + +subsection {* Other simple lemmas *} + declare disj_absorb [simp] conj_absorb [simp] lemma ex1_eq[iff]: "EX! x. x = t" "EX! x. t = x" by blast+ + theorem choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" apply (rule iffI) apply (rule_tac a = "%x. THE y. P x y" in ex1I) @@ -1142,7 +1156,7 @@ by(rule trans[OF trans[OF c a] arg_cong[OF c, of "f y"]]) -subsubsection {* Generic cases and induction *} +subsection {* Generic cases and induction *} constdefs induct_forall :: "('a => bool) => bool" diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Hyperreal/Integration.thy --- a/src/HOL/Hyperreal/Integration.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Hyperreal/Integration.thy Tue Feb 01 18:01:57 2005 +0100 @@ -263,12 +263,14 @@ apply (drule partition_lhs, auto) apply (simp split: nat_diff_split) apply (subst partition) -apply (subst lemma_partition_append2, assumption+) +apply (simplesubst lemma_partition_append2, assumption+) + --{*Need to substitute the last occurrence*} apply (rule conjI) apply (drule_tac [2] lemma_partition_append1) apply (auto simp add: partition_lhs partition_rhs) done + text{*We can always find a division that is fine wrt any gauge*} lemma partition_exists: diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Hyperreal/MacLaurin.thy --- a/src/HOL/Hyperreal/MacLaurin.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Hyperreal/MacLaurin.thy Tue Feb 01 18:01:57 2005 +0100 @@ -579,7 +579,7 @@ in Maclaurin_all_le_objl) apply safe apply simp - apply (subst mod_Suc_eq_Suc_mod) + apply (simplesubst mod_Suc_eq_Suc_mod) --{*simultaneous substitution*} apply (cut_tac m=m in mod_exhaust_less_4, safe, simp+) apply (rule DERIV_minus, simp+) apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Hyperreal/Transcendental.thy --- a/src/HOL/Hyperreal/Transcendental.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Hyperreal/Transcendental.thy Tue Feb 01 18:01:57 2005 +0100 @@ -138,8 +138,7 @@ lemma real_pow_sqrt_eq_sqrt_pow: "0 \ x ==> (sqrt x)\ = sqrt(x\)" apply (simp add: sqrt_def) -apply (subst numeral_2_eq_2) -apply (auto intro: real_root_pow_pos2 [THEN ssubst] real_root_pos2 [THEN ssubst] simp del: realpow_Suc) +apply (simp only: numeral_2_eq_2 real_root_pow_pos2 real_root_pos2) done lemma real_pow_sqrt_eq_sqrt_abs_pow2: @@ -1424,7 +1423,9 @@ apply (drule sums_minus) apply (rule neg_less_iff_less [THEN iffD1]) apply (frule sums_unique, auto simp add: times_divide_eq) -apply (rule_tac y = "sumr 0 (Suc (Suc (Suc 0))) (%n. - ((- 1) ^ n / (real (fact (2 * n))) * 2 ^ (2 * n)))" in order_less_trans) +apply (rule_tac y = "sumr 0 (Suc (Suc (Suc 0))) + (%n. - ((- 1) ^ n / (real(fact (2*n))) * 2 ^ (2*n)))" + in order_less_trans) apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc realpow_Suc) apply (simp (no_asm) add: mult_assoc times_divide_eq del: sumr_Suc) apply (rule sumr_pos_lt_pair) @@ -1435,12 +1436,11 @@ apply (rule real_of_nat_fact_gt_zero)+ apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc) apply (subst fact_lemma) -apply (subst fact_Suc) -apply (subst real_of_nat_mult) -apply (erule ssubst, subst real_of_nat_mult) +apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"]) +apply (simp only: real_of_nat_mult) apply (rule real_mult_less_mono, force) -prefer 2 apply force -apply (rule_tac [2] real_of_nat_fact_gt_zero) + apply (rule_tac [3] real_of_nat_fact_gt_zero) + prefer 2 apply force apply (rule real_of_nat_less_iff [THEN iffD2]) apply (rule fact_less_mono, auto) done diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/IMP/Denotation.thy --- a/src/HOL/IMP/Denotation.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/IMP/Denotation.thy Tue Feb 01 18:01:57 2005 +0100 @@ -34,15 +34,13 @@ lemma C_While_If: "C(\ b \ c) = C(\ b \ c;\ b \ c \ \)" apply (simp (no_asm)) -apply (subst lfp_unfold [OF Gamma_mono]) back back -apply (subst Gamma_def [THEN meta_eq_to_obj_eq, THEN fun_cong]) -apply simp +apply (subst lfp_unfold [OF Gamma_mono]) --{*lhs only*} +apply (simp add: Gamma_def) done (* Operational Semantics implies Denotational Semantics *) lemma com1: "\c,s\ \\<^sub>c t \ (s,t) \ C(c)" - (* start with rule induction *) apply (erule evalc.induct) apply auto diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Induct/LList.thy --- a/src/HOL/Induct/LList.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Induct/LList.thy Tue Feb 01 18:01:57 2005 +0100 @@ -673,12 +673,9 @@ subsection{* Deriving @{text llist_equalityI} -- @{text llist} equality is a bisimulation *} lemma LListD_Fun_subset_Times_llist: - "r <= (llist A) <*> (llist A) ==> - LListD_Fun (diag A) r <= (llist A) <*> (llist A)" -apply (unfold LListD_Fun_def) -apply (subst llist_unfold) -apply (simp add: NIL_def CONS_def, fast) -done + "r <= (llist A) <*> (llist A) + ==> LListD_Fun (diag A) r <= (llist A) <*> (llist A)" +by (auto simp add: LListD_Fun_def) lemma subset_Times_llist: "prod_fun Rep_LList Rep_LList ` r <= @@ -866,8 +863,9 @@ lemma lappend_iterates [simp]: "lappend (iterates f x) N = iterates f x" apply (rule_tac r = "range (%u. (lappend (iterates f u) N,iterates f u))" in llist_equalityI) -apply (rule rangeI, safe) -apply (subst iterates, simp) + apply (rule rangeI) +apply (safe) +apply (simplesubst iterates, simp) --{*two redexes*} done subsubsection{* Two proofs that @{text lmap} distributes over lappend *} diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Integ/IntDef.thy --- a/src/HOL/Integ/IntDef.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Integ/IntDef.thy Tue Feb 01 18:01:57 2005 +0100 @@ -849,7 +849,7 @@ done lemma int_setsum: "int (setsum f A) = setsum (int \ f) A" - by (subst int_eq_of_nat, rule setsum_of_nat) + by (simp add: int_eq_of_nat setsum_of_nat) lemma setprod_of_nat: "of_nat (setprod f A) = setprod (of_nat \ f) A" apply (case_tac "finite A") @@ -862,7 +862,7 @@ done lemma int_setprod: "int (setprod f A) = setprod (int \ f) A" - by (subst int_eq_of_nat, rule setprod_of_nat) + by (simp add: int_eq_of_nat setprod_of_nat) lemma setsum_constant [simp]: "finite A ==> (\x \ A. y) = of_nat(card A) * y" apply (erule finite_induct) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/IsaMakefile Tue Feb 01 18:01:57 2005 +0100 @@ -74,13 +74,14 @@ $(SRC)/Provers/Arith/fast_lin_arith.ML $(SRC)/Provers/blast.ML \ $(SRC)/Provers/clasimp.ML $(SRC)/Provers/classical.ML \ $(SRC)/Provers/hypsubst.ML $(SRC)/Provers/induct_method.ML \ - $(SRC)/Provers/make_elim.ML $(SRC)/Provers/order.ML \ - $(SRC)/Provers/quasi.ML \ + $(SRC)/Provers/make_elim.ML $(SRC)/Provers/order.ML $(SRC)/Provers/quasi.ML\ $(SRC)/Provers/simplifier.ML $(SRC)/Provers/splitter.ML \ $(SRC)/Provers/trancl.ML \ $(SRC)/TFL/dcterm.ML $(SRC)/TFL/post.ML $(SRC)/TFL/rules.ML \ $(SRC)/TFL/tfl.ML $(SRC)/TFL/thms.ML $(SRC)/TFL/thry.ML \ $(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML \ + $(SRC)/Provers/eqsubst.ML\ + eqrule_HOL_data.ML\ Datatype.thy Datatype_Universe.thy \ Divides.thy Equiv_Relations.thy Extraction.thy Finite_Set.ML Finite_Set.thy \ Fun.thy Gfp.thy Hilbert_Choice.thy HOL.ML \ diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Library/Word.thy Tue Feb 01 18:01:57 2005 +0100 @@ -455,7 +455,7 @@ proof (rule n_div_2_cases [of n]) assume [simp]: "n = 0" show ?thesis - apply (subst nat_to_bv_helper.simps [of n]) + apply (simp only: nat_to_bv_helper.simps [of n]) apply simp done next @@ -467,7 +467,7 @@ have ind': "\l. nat_to_bv_helper (n div 2) l = nat_to_bv_helper (n div 2) [] @ l" by blast show ?thesis - apply (subst nat_to_bv_helper.simps [of n]) + apply (simp only: nat_to_bv_helper.simps [of n]) apply (cases "n=0") apply simp apply (simp only: if_False) @@ -491,7 +491,7 @@ assume [simp]: "0 < n" show ?thesis apply (subst nat_to_bv_def [of n]) - apply (subst nat_to_bv_helper.simps [of n]) + apply (simp only: nat_to_bv_helper.simps [of n]) apply (subst unfold_nat_to_bv_helper) using prems apply simp @@ -545,7 +545,7 @@ by simp show ?thesis apply (subst nat_to_bv_def) - apply (subst nat_to_bv_helper.simps [of n]) + apply (simp only: nat_to_bv_helper.simps [of n]) apply (simp only: n0' if_False) apply (subst unfold_nat_to_bv_helper) apply (subst bv_to_nat_dist_append) @@ -2425,7 +2425,7 @@ from ki ij jl lw show ?thesis - apply (subst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"]) + apply (simplesubst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"]) apply simp_all apply (rule w_def) apply (simp add: w_defs min_def) @@ -2566,7 +2566,7 @@ apply (subst indq [symmetric]) apply (subst indq [symmetric]) apply auto - apply (subst nat_to_bv_helper.simps [of "2 * qq + 1"]) + apply (simp only: nat_to_bv_helper.simps [of "2 * qq + 1"]) apply simp apply safe apply (subgoal_tac "2 * qq + 1 ~= 2 * q") @@ -2575,7 +2575,7 @@ apply (subgoal_tac "(2 * qq + 1) div 2 = qq") apply simp apply (subst zdiv_zadd1_eq,simp) - apply (subst nat_to_bv_helper.simps [of "2 * qq"]) + apply (simp only: nat_to_bv_helper.simps [of "2 * qq"]) apply simp done qed diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Matrix/Matrix.thy --- a/src/HOL/Matrix/Matrix.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Matrix/Matrix.thy Tue Feb 01 18:01:57 2005 +0100 @@ -129,7 +129,7 @@ lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)" apply (simp add: one_matrix_def) -apply (subst RepAbs_matrix) +apply (simplesubst RepAbs_matrix) apply (rule exI[of _ n], simp add: split_if)+ by (simp add: split_if, arith) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Matrix/MatrixGeneral.thy --- a/src/HOL/Matrix/MatrixGeneral.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Matrix/MatrixGeneral.thy Tue Feb 01 18:01:57 2005 +0100 @@ -699,16 +699,16 @@ let ?mx = "max (ncols a) (max (nrows u) (nrows v))" from prems show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) = combine_matrix fadd (mult_matrix_n (max (ncols a) (nrows u)) fmul fadd a u) (mult_matrix_n (max (ncols a) (nrows v)) fmul fadd a v)" - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simplesubst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simplesubst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simplesubst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd]) apply (simp add: combine_matrix_def combine_infmatrix_def) apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) - apply (subst RepAbs_matrix) + apply (simplesubst RepAbs_matrix) apply (simp, auto) apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero) @@ -739,16 +739,16 @@ let ?mx = "max (nrows a) (max (ncols u) (ncols v))" from prems show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a = combine_matrix fadd (mult_matrix_n (max (ncols u) (nrows a)) fmul fadd u a) (mult_matrix_n (max (ncols v) (nrows a)) fmul fadd v a)" - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simplesubst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simplesubst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) apply (simp add: max1 max2 combine_nrows combine_ncols)+ apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd]) apply (simp add: combine_matrix_def combine_infmatrix_def) apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) - apply (subst RepAbs_matrix) + apply (simplesubst RepAbs_matrix) apply (simp, auto) apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero) apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) @@ -792,18 +792,19 @@ lemma mult_matrix_zero_closed: "\fadd 0 0 = 0; zero_closed fmul\ \ zero_closed (mult_matrix fmul fadd)" apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def) - apply (auto) - by (subst foldseq_zero, (simp add: zero_matrix_def)+)+ + apply (simp add: foldseq_zero zero_matrix_def) + done + lemma mult_matrix_n_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix_n n fmul fadd A 0 = 0" apply (simp add: mult_matrix_n_def) - apply (subst foldseq_zero) - by (simp_all add: zero_matrix_def) + apply (simp add: foldseq_zero zero_matrix_def) + done lemma mult_matrix_n_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix_n n fmul fadd 0 A = 0" apply (simp add: mult_matrix_n_def) - apply (subst foldseq_zero) - by (simp_all add: zero_matrix_def) + apply (simp add: foldseq_zero zero_matrix_def) + done lemma mult_matrix_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix fmul fadd 0 A = 0" by (simp add: mult_matrix_def) @@ -851,7 +852,7 @@ apply (subst RepAbs_matrix) apply (rule exI[of _ "Suc m"], simp) apply (rule exI[of _ "Suc n"], simp+) -by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+ +by (simplesubst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+ lemma apply_singleton_matrix[simp]: "f 0 = 0 \ apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))" apply (subst Rep_matrix_inject[symmetric]) @@ -892,7 +893,7 @@ lemma combine_singleton: "f 0 0 = 0 \ combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)" apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def) -apply (subst RepAbs_matrix) +apply (simplesubst RepAbs_matrix) apply (rule exI[of _ "Suc j"], simp) apply (rule exI[of _ "Suc i"], simp) apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) @@ -911,7 +912,7 @@ (if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))" apply (simp add: move_matrix_def) apply (auto) -by (subst RepAbs_matrix, +by (simplesubst RepAbs_matrix, rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith, rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+ @@ -946,7 +947,7 @@ "Rep_matrix (take_columns A c) j i = (if i < c then (Rep_matrix A j i) else 0)" apply (simp add: take_columns_def) -apply (subst RepAbs_matrix) +apply (simplesubst RepAbs_matrix) apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) done @@ -955,7 +956,7 @@ "Rep_matrix (take_rows A r) j i = (if j < r then (Rep_matrix A j i) else 0)" apply (simp add: take_rows_def) -apply (subst RepAbs_matrix) +apply (simplesubst RepAbs_matrix) apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) done @@ -1163,21 +1164,21 @@ apply (simp add: Rep_matrix_inject[THEN sym]) apply (rule ext)+ apply (simp add: mult_matrix_def) - apply (subst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"]) + apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ - apply (subst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ - apply (subst mult_matrix_nm[of _ _ _ "?N"]) + apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ - apply (subst mult_matrix_nm[of _ _ _ "?N"]) + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ - apply (subst mult_matrix_nm[of _ _ _ "?N"]) + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ - apply (subst mult_matrix_nm[of _ _ _ "?N"]) + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ apply (simp add: mult_matrix_n_def) apply (rule comb_left) apply ((rule ext)+, simp) - apply (subst RepAbs_matrix) + apply (simplesubst RepAbs_matrix) apply (rule exI[of _ "nrows B"]) apply (simp add: nrows prems foldseq_zero) apply (rule exI[of _ "ncols C"]) @@ -1227,7 +1228,7 @@ have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)" by (simp! add: associative_def commutative_def) then show ?concl - apply (subst mult_matrix_assoc) + apply (simplesubst mult_matrix_assoc) apply (simp_all!) by (simp add: associative_def distributive_def l_distributive_def r_distributive_def)+ qed @@ -1361,7 +1362,7 @@ "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B" apply (simp! add: le_matrix_def Rep_mult_matrix) apply (auto) - apply (subst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ + apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ apply (rule le_foldseq) by (auto) @@ -1378,7 +1379,7 @@ "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C" apply (simp! add: le_matrix_def Rep_mult_matrix) apply (auto) - apply (subst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ + apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ apply (rule le_foldseq) by (auto) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Matrix/SparseMatrix.thy --- a/src/HOL/Matrix/SparseMatrix.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Matrix/SparseMatrix.thy Tue Feb 01 18:01:57 2005 +0100 @@ -271,7 +271,7 @@ apply (intro strip) apply (frule_tac as=brr in sorted_spvec_cons1) apply (simp add: ring_eq_simps sparse_row_matrix_cons) - apply (subst Rep_matrix_zero_imp_mult_zero) + apply (simplesubst Rep_matrix_zero_imp_mult_zero) apply (simp) apply (intro strip) apply (rule disjI2) @@ -805,9 +805,9 @@ apply auto apply (case_tac "x=a") apply (simp) - apply (subst sorted_sparse_row_matrix_zero) + apply (simplesubst sorted_sparse_row_matrix_zero) apply auto - apply (subst Rep_sparse_row_vector_zero) + apply (simplesubst Rep_sparse_row_vector_zero) apply (simp_all add: neg_def) done diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/BV/Effect.thy --- a/src/HOL/MicroJava/BV/Effect.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/BV/Effect.thy Tue Feb 01 18:01:57 2005 +0100 @@ -6,7 +6,10 @@ header {* \isaheader{Effect of Instructions on the State Type} *} -theory Effect = JVMType + JVMExceptions: +theory Effect +imports JVMType "../JVM/JVMExceptions" +begin + types succ_type = "(p_count \ state_type option) list" diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Tue Feb 01 18:01:57 2005 +0100 @@ -295,7 +295,7 @@ apply simp apply assumption apply simp apply assumption apply simp -apply (subst method_rec) apply simp +apply (simplesubst method_rec) apply simp apply force apply simp apply (simp only: map_add_def) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/Comp/CorrCompTp.thy --- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Tue Feb 01 18:01:57 2005 +0100 @@ -3,7 +3,9 @@ Author: Martin Strecker *) -theory CorrCompTp = LemmasComp + JVM + TypeInf + Altern: +theory CorrCompTp +imports LemmasComp TypeInf "../BV/JVM" "../BV/Altern" +begin declare split_paired_All [simp del] declare split_paired_Ex [simp del] diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/Comp/DefsComp.thy --- a/src/HOL/MicroJava/Comp/DefsComp.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/Comp/DefsComp.thy Tue Feb 01 18:01:57 2005 +0100 @@ -5,8 +5,9 @@ (* Definitions for accessing parts of methods, states etc. *) -theory DefsComp = JVMExec: - +theory DefsComp +imports "../JVM/JVMExec" +begin constdefs diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/Comp/TranslCompTp.thy --- a/src/HOL/MicroJava/Comp/TranslCompTp.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy Tue Feb 01 18:01:57 2005 +0100 @@ -3,7 +3,10 @@ Author: Martin Strecker *) -theory TranslCompTp = JVMType + Index: +theory TranslCompTp +imports Index "../BV/JVMType" +begin + (**********************************************************************) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/MicroJava/J/WellForm.thy --- a/src/HOL/MicroJava/J/WellForm.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/MicroJava/J/WellForm.thy Tue Feb 01 18:01:57 2005 +0100 @@ -542,17 +542,13 @@ apply hypsubst apply simp apply (erule conjE) -apply (subst method_rec) - apply (assumption) - apply (assumption) +apply (simplesubst method_rec, assumption+) apply (clarify) apply (erule_tac x = "Da" in allE) apply (clarsimp) apply (simp add: map_of_map) apply (clarify) - apply (subst method_rec) - apply (assumption) - apply (assumption) + apply (subst method_rec, assumption+) apply (simp add: map_add_def map_of_map split add: option.split) done @@ -566,17 +562,13 @@ apply hypsubst apply simp apply (erule conjE) -apply (subst method_rec) - apply (assumption) - apply (assumption) +apply (simplesubst method_rec, assumption+) apply (clarify) apply (erule_tac x = "Da" in allE) apply (clarsimp) apply (simp add: map_of_map) apply (clarify) - apply (subst method_rec) - apply (assumption) - apply (assumption) + apply (subst method_rec, assumption+) apply (simp add: map_add_def map_of_map split add: option.split) done diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/NumberTheory/EulerFermat.thy --- a/src/HOL/NumberTheory/EulerFermat.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/NumberTheory/EulerFermat.thy Tue Feb 01 18:01:57 2005 +0100 @@ -259,7 +259,7 @@ \(BnorRset(a, m)) * x^card (BnorRset (a, m))" apply (induct a m rule: BnorRset_induct) prefer 2 - apply (subst BnorRset.simps) + apply (simplesubst BnorRset.simps) --{*multiple redexes*} apply (unfold Let_def, auto) apply (simp add: Bnor_fin Bnor_mem_zle_swap) apply (subst setprod_insert) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/OrderedGroup.thy Tue Feb 01 18:01:57 2005 +0100 @@ -435,14 +435,14 @@ qed lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::pordered_ab_group_add))" -apply (subst less_iff_diff_less_0) +apply (subst less_iff_diff_less_0 [of a]) apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst]) apply (simp add: diff_minus add_ac) done lemma less_diff_eq: "(a < c-b) = (a + (b::'a::pordered_ab_group_add) < c)" -apply (subst less_iff_diff_less_0) -apply (rule less_iff_diff_less_0 [of _ "c-b", THEN ssubst]) +apply (subst less_iff_diff_less_0 [of "a+b"]) +apply (subst less_iff_diff_less_0 [of a]) apply (simp add: diff_minus add_ac) done @@ -608,8 +608,8 @@ lemma meet_0_imp_0: "meet a (-a) = 0 \ a = (0::'a::lordered_ab_group)" apply (simp add: meet_eq_neg_join) apply (simp add: join_comm) -apply (subst join_0_imp_0) -by auto +apply (erule join_0_imp_0) +done lemma join_0_eq_0[simp]: "(join a (-a) = 0) = (a = (0::'a::lordered_ab_group))" by (auto, erule join_0_imp_0) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Product_Type.thy Tue Feb 01 18:01:57 2005 +0100 @@ -451,9 +451,7 @@ lemma split_split: "R (split c p) = (ALL x y. p = (x, y) --> R (c x y))" -- {* For use with @{text split} and the Simplifier. *} - apply (subst surjective_pairing) - apply (subst split_conv, blast) - done + by (insert surj_pair [of p], clarify, simp) text {* @{thm [source] split_split} could be declared as @{text "[split]"} diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Recdef.thy --- a/src/HOL/Recdef.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Recdef.thy Tue Feb 01 18:01:57 2005 +0100 @@ -8,7 +8,6 @@ theory Recdef imports Wellfounded_Relations Datatype files - ("../TFL/isand.ML") ("../TFL/casesplit.ML") ("../TFL/utils.ML") ("../TFL/usyntax.ML") @@ -45,7 +44,6 @@ lemma tfl_exE: "\x. P x ==> \x. P x --> Q ==> Q" by blast -use "../TFL/isand.ML" use "../TFL/casesplit.ML" use "../TFL/utils.ML" use "../TFL/usyntax.ML" diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/Ring_and_Field.thy Tue Feb 01 18:01:57 2005 +0100 @@ -1546,7 +1546,9 @@ by (simp only: abs_prts[of a] abs_prts[of b] ring_eq_simps) { fix u v :: 'a - have bh: "\u = a; v = b\ \ u * v = ?y" + have bh: "\u = a; v = b\ \ + u * v = pprt a * pprt b + pprt a * nprt b + + nprt a * pprt b + nprt a * nprt b" apply (subst prts[of u], subst prts[of v]) apply (simp add: left_distrib right_distrib add_ac) done diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/SET-Protocol/Purchase.thy --- a/src/HOL/SET-Protocol/Purchase.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/SET-Protocol/Purchase.thy Tue Feb 01 18:01:57 2005 +0100 @@ -51,11 +51,11 @@ consts CardSecret :: "nat => nat" - (*Maps Cardholders to CardSecrets. - A CardSecret of 0 means no cerificate, must use unsigned format.*) + --{*Maps Cardholders to CardSecrets. + A CardSecret of 0 means no cerificate, must use unsigned format.*} PANSecret :: "nat => nat" - (*Maps Cardholders to PANSecrets.*) + --{*Maps Cardholders to PANSecrets.*} set_pur :: "event list set" @@ -372,7 +372,7 @@ apply (auto simp add: used_Cons symKeys_neq_imp_neq) done -(*General facts about message reception*) +text{*General facts about message reception*} lemma Gets_imp_Says: "[| Gets B X \ set evs; evs \ set_pur |] ==> \A. Says A B X \ set evs" @@ -425,30 +425,28 @@ evs \ set_pur|] ==> priEK C \ KK | C \ bad" by auto -text{*trivial proof because (priEK C) never appears even in (parts evs)*} +text{*trivial proof because @{term"priEK C"} never appears even in + @{term "parts evs"}. *} lemma analz_image_priEK: "evs \ set_pur ==> (Key (priEK C) \ analz (Key`KK Un (knows Spy evs))) = (priEK C \ KK | C \ bad)" by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) + subsection{*Public Keys in Certificates are Correct*} lemma Crypt_valid_pubEK [dest!]: "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|} \ parts (knows Spy evs); evs \ set_pur |] ==> EKi = pubEK C" -apply (erule rev_mp) -apply (erule set_pur.induct, auto) -done +by (erule rev_mp, erule set_pur.induct, auto) lemma Crypt_valid_pubSK [dest!]: "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|} \ parts (knows Spy evs); evs \ set_pur |] ==> SKi = pubSK C" -apply (erule rev_mp) -apply (erule set_pur.induct, auto) -done +by (erule rev_mp, erule set_pur.induct, auto) lemma certificate_valid_pubEK: "[| cert C EKi onlyEnc (priSK RCA) \ parts (knows Spy evs); @@ -730,9 +728,7 @@ lemma M_Notes_PG: "[|Notes M {|Number LID_M, Agent P, Agent M, Agent C, etc|} \ set evs; evs \ set_pur|] ==> \j. P = PG j" -apply (erule rev_mp) -apply (erule set_pur.induct, simp_all) -done +by (erule rev_mp, erule set_pur.induct, simp_all) text{*If we trust M, then @{term LID_M} determines his choice of P (Payment Gateway)*} @@ -816,7 +812,7 @@ done text{*When M receives AuthRes, he knows that P signed it, including - the identifying tages and the purchase amount, which he can verify. + the identifying tags and the purchase amount, which he can verify. (Although the spec has SIGNED and UNSIGNED forms of AuthRes, they send the same message to M.)*} theorem M_verifies_AuthRes: @@ -1015,9 +1011,7 @@ "[|Notes (Cardholder k) {|Number n, Agent P, Agent (Cardholder k), Agent C, etc|} \ set evs; evs \ set_pur|] ==> False" -apply (erule rev_mp) -apply (erule set_pur.induct, auto) -done +by (erule rev_mp, erule set_pur.induct, auto) text{*When M sees a dual signature, he knows that it originated with C. Using XID he knows it was intended for him. @@ -1166,5 +1160,3 @@ done end - - diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/UNITY/FP.thy --- a/src/HOL/UNITY/FP.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/UNITY/FP.thy Tue Feb 01 18:01:57 2005 +0100 @@ -19,20 +19,19 @@ "FP F == {s. F : stable {s}}" lemma stable_FP_Orig_Int: "F : stable (FP_Orig F Int B)" -apply (unfold FP_Orig_def stable_def) -apply (subst Int_Union2) +apply (simp only: FP_Orig_def stable_def Int_Union2) apply (blast intro: constrains_UN) done lemma FP_Orig_weakest: "(!!B. F : stable (A Int B)) ==> A <= FP_Orig F" -by (unfold FP_Orig_def stable_def, blast) +by (simp add: FP_Orig_def stable_def, blast) lemma stable_FP_Int: "F : stable (FP F Int B)" apply (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x}) ") prefer 2 apply blast apply (simp (no_asm_simp) add: Int_insert_right) -apply (unfold FP_def stable_def) +apply (simp add: FP_def stable_def) apply (rule constrains_UN) apply (simp (no_asm)) done @@ -40,7 +39,7 @@ lemma FP_equivalence: "FP F = FP_Orig F" apply (rule equalityI) apply (rule stable_FP_Int [THEN FP_Orig_weakest]) -apply (unfold FP_Orig_def FP_def, clarify) +apply (simp add: FP_Orig_def FP_def, clarify) apply (drule_tac x = "{x}" in spec) apply (simp add: Int_insert_right) done diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/UNITY/ListOrder.thy --- a/src/HOL/UNITY/ListOrder.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/UNITY/ListOrder.thy Tue Feb 01 18:01:57 2005 +0100 @@ -215,8 +215,7 @@ apply (erule disjE) apply (simp_all (no_asm_simp) add: neq_Nil_conv nth_append) apply (blast intro: genPrefix.append, auto) -apply (subst append_cons_eq) -apply (blast intro: genPrefix_append_both genPrefix.append) +apply (subst append_cons_eq, fast intro: genPrefix_append_both genPrefix.append) done lemma append_one_genPrefix: diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/UNITY/Rename.thy --- a/src/HOL/UNITY/Rename.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/UNITY/Rename.thy Tue Feb 01 18:01:57 2005 +0100 @@ -110,8 +110,8 @@ lemma bij_extend: "bij h ==> bij (extend (%(x,u::'c). h x))" apply (rule bijI) apply (blast intro: Extend.inj_extend) -apply (rule_tac f = "extend (% (x,u) . inv h x) " in surjI) -apply (subst inv_inv_eq [of h, symmetric], assumption) +apply (rule_tac f = "extend (% (x,u) . inv h x) " in surjI) +apply (simplesubst inv_inv_eq [of h, symmetric], assumption) apply (subst extend_inv, simp add: bij_imp_bij_inv) apply (simp add: inv_inv_eq) apply (rule Extend.extend_inverse) diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/eqrule_HOL_data.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/eqrule_HOL_data.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,68 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/eqrule_HOL_data.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Modified: 22 July 2004 + Created: 18 Feb 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Data for equality rules in the logic + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +structure HOL_EqRuleData : EQRULE_DATA = +struct + +val eq_reflection = thm "eq_reflection"; +val mp = thm "mp"; +val spec = thm "spec"; +val if_bool_eq_conj = thm "if_bool_eq_conj"; +val iffD1 = thm "iffD1"; +val conjunct2 = thm "conjunct2"; +val conjunct1 = thm "conjunct1"; + +fun mk_eq th = case concl_of th of + Const("==",_)$_$_ => Some (th) + | _$(Const("op =",_)$_$_) => Some (th RS eq_reflection) + | _ => None; + +val tranformation_pairs = + [("op -->", [mp]), ("op &", [conjunct1,conjunct2]), + ("All", [spec]), ("True", []), ("False", []), + ("If", [if_bool_eq_conj RS iffD1])]; + +(* +val mk_atomize: (string * thm list) list -> thm -> thm list +looks too specific to move it somewhere else +*) +fun mk_atomize pairs = + let fun atoms th = + (case Thm.concl_of th of + Const("Trueprop",_) $ p => + (case Term.head_of p of + Const(a,_) => + (case Library.assoc(pairs,a) of + Some(rls) => flat (map atoms ([th] RL rls)) + | None => [th]) + | _ => [th]) + | _ => [th]) + in atoms end; + +val prep_meta_eq = + (mapfilter + mk_eq + o (mk_atomize tranformation_pairs) + o Drule.gen_all + o zero_var_indexes) + +end; + +structure EqRuleData = HOL_EqRuleData; + +structure EQSubstTac = + EQSubstTacFUN(structure EqRuleData = EqRuleData); + + diff -r cb3612cc41a3 -r fc075ae929e4 src/HOL/ex/set.thy --- a/src/HOL/ex/set.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/HOL/ex/set.thy Tue Feb 01 18:01:57 2005 +0100 @@ -110,9 +110,9 @@ apply (rule_tac [2] inj_on_inv) apply (erule subset_inj_on [OF _ subset_UNIV]) txt {* Tricky variable instantiations! *} - apply (erule ssubst, subst double_complement) + apply (erule ssubst, simplesubst double_complement) apply (rule subsetI, erule imageE, erule ssubst, rule rangeI) - apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric]) + apply (erule ssubst, simplesubst double_complement, erule inv_image_comp [symmetric]) done diff -r cb3612cc41a3 -r fc075ae929e4 src/Provers/eqsubst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/eqsubst.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,212 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/eqsubst_tac.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Created: 29 Jan 2005 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + A Tactic to perform a substiution using an equation. + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +(* Logic specific data *) +signature EQRULE_DATA = +sig + (* to make a meta equality theorem in the current logic *) + val prep_meta_eq : thm -> thm list +end; + +(* the signature of an instance of the SQSUBST tactic *) +signature EQSUBST_TAC = +sig + val eqsubst_asm_meth : Thm.thm list -> Proof.method + val eqsubst_asm_tac : Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_asm_tac' : Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_meth : Thm.thm list -> Proof.method + val eqsubst_tac : Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_tac' : Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq + val meth : bool * Thm.thm list -> Proof.context -> Proof.method + val subst : Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq + val subst_asm : Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq + + val setup : (Theory.theory -> Theory.theory) list +end; + +functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA) +(* : EQSUBST_TAC *) += struct + +fun search_tb_lr_f f ft = + let + fun maux ft = + let val t' = (IsaFTerm.focus_of_fcterm ft) + (* val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *) + in + (case t' of + (_ $ _) => Seq.append(f ft, + Seq.append(maux (IsaFTerm.focus_left ft), + maux (IsaFTerm.focus_right ft))) + | (Abs _) => Seq.append (f ft, maux (IsaFTerm.focus_abs ft)) + | leaf => f ft) end + in maux ft end; + +fun search_for_match sgn lhs maxidx = + IsaFTerm.find_fcterm_matches + search_tb_lr_f + (IsaFTerm.clean_unify_ft sgn maxidx lhs); + + +(* CLEANUP: lots of duplication of code for substituting in +assumptions and conclusion - this could be cleaned up a little. *) + +fun subst_concl rule cfvs i th (conclthm, concl_matches)= + let + fun apply_subst m = + (RWInst.rw m rule conclthm) + |> IsaND.schemify_frees_to_vars cfvs + |> RWInst.beta_eta_contract_tac + |> (fn r => Tactic.rtac r i th) + |> Seq.map Drule.zero_var_indexes + in + Seq.flat (Seq.map apply_subst concl_matches) + end; + + +(* substitute within the conclusion of goal i of gth, using a meta +equation rule *) +fun subst rule i gth = + let + val th = Thm.incr_indexes 1 gth; + val tgt_term = Thm.prop_of th; + val maxidx = Term.maxidx_of_term tgt_term; + + val rule' = Drule.zero_var_indexes rule; + val (lhs,_) = Logic.dest_equals (Thm.concl_of rule'); + + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + val trivify = Thm.trivial o ctermify; + + val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; + val cfvs = rev (map ctermify fvs); + + val conclthm = trivify (Logic.strip_imp_concl fixedbody); + val concl_matches = + search_for_match sgn lhs maxidx + ((IsaFTerm.focus_right + o IsaFTerm.focus_left + o IsaFTerm.fcterm_of_term + o Thm.prop_of) conclthm); + in + subst_concl rule' cfvs i th (conclthm, concl_matches) + end; + +(* substitute using an object or meta level equality *) +fun eqsubst_tac' instepthm i th = + let val stepthms = Seq.of_list (EqRuleData.prep_meta_eq instepthm) in + Seq.flat (Seq.map (fn rule => subst rule i th) stepthms) + end; +(* substitute using one of the given theorems *) +fun eqsubst_tac instepthms i th = + Seq.flat (Seq.map (fn r => eqsubst_tac' r i th) (Seq.of_list instepthms)); + +(* inthms are the given arguments in Isar, and treated as eqstep with + the first one, then the second etc *) +fun eqsubst_meth inthms = + Method.METHOD + (fn facts => + HEADGOAL (eqsubst_tac inthms THEN' Method.insert_tac facts)); + + +fun apply_subst_in_asm rule cfvs i th matchseq = + let + fun apply_subst ((j, pth), mseq) = + Seq.flat (Seq.map + (fn m => + (RWInst.rw m rule pth) + |> Thm.permute_prems 0 ~1 + |> IsaND.schemify_frees_to_vars cfvs + |> RWInst.beta_eta_contract_tac + |> (fn r => Tactic.dtac r i th) + |> Seq.map Drule.zero_var_indexes) + mseq) + in + Seq.flat (Seq.map apply_subst matchseq) + end; + + +(* substitute within an assumption of goal i of gth, using a meta +equation rule *) +fun subst_asm rule i gth = + let + val th = Thm.incr_indexes 1 gth; + val tgt_term = Thm.prop_of th; + val maxidx = Term.maxidx_of_term tgt_term; + + val rule' = Drule.zero_var_indexes rule; + val (lhs,_) = Logic.dest_equals (Thm.concl_of rule'); + + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + val trivify = Thm.trivial o ctermify; + + val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; + val cfvs = rev (map ctermify fvs); + + val premthms = Seq.of_list (IsaPLib.number_list 1 + (map trivify (Logic.strip_imp_prems fixedbody))); + val prem_matches = + Seq.map (fn (i, pth) => + ((i, pth), search_for_match sgn lhs maxidx + ((IsaFTerm.focus_right + o IsaFTerm.fcterm_of_term + o Thm.prop_of) pth))) + premthms; + in + apply_subst_in_asm rule' cfvs i th prem_matches + end; + +(* substitute using an object or meta level equality *) +fun eqsubst_asm_tac' instepthm i th = + let val stepthms = Seq.of_list (EqRuleData.prep_meta_eq instepthm) in + Seq.flat (Seq.map (fn rule => subst_asm rule i th) stepthms) + end; + +(* substitute using one of the given theorems *) +fun eqsubst_asm_tac instepthms i th = + Seq.flat (Seq.map (fn r => eqsubst_asm_tac' r i th) + (Seq.of_list instepthms)); + +(* inthms are the given arguments in Isar, and treated as eqstep with + the first one, then the second etc *) +fun eqsubst_asm_meth inthms = + Method.METHOD + (fn facts => + HEADGOAL (eqsubst_asm_tac inthms THEN' Method.insert_tac facts)); + +(* combination method that takes a flag (true indicates that subst +should be done to an assumption, false = apply to the conclusion of +the goal) as well as the theorems to use *) +fun meth (asmflag, inthms) ctxt = + if asmflag then eqsubst_asm_meth inthms else eqsubst_meth inthms; + +(* syntax for options, given "(asm)" will give back true, without + gives back false *) +val options_syntax = + (Args.parens (Args.$$$ "asm") >> (K true)) || + (Scan.succeed false); + +(* method syntax, first take options, then theorems *) +fun meth_syntax meth src ctxt = + meth (snd (Method.syntax ((Scan.lift options_syntax) + -- Attrib.local_thms) src ctxt)) + ctxt; + +(* setup function for adding method to theory. *) +val setup = + [Method.add_method ("subst", meth_syntax meth, "Substiution with an equation. Use \"(asm)\" option to substitute in an assumption.")]; + +end; \ No newline at end of file diff -r cb3612cc41a3 -r fc075ae929e4 src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML Sun Jan 30 20:48:50 2005 +0100 +++ b/src/Provers/hypsubst.ML Tue Feb 01 18:01:57 2005 +0100 @@ -247,6 +247,6 @@ val hypsubst_setup = [Method.add_methods [("hypsubst", hyp_subst_meth, "substitution using an assumption (improper)"), - ("subst", subst_meth, "substitution")]]; + ("simplesubst", subst_meth, "simple substitution")]]; end; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Sun Jan 30 20:48:50 2005 +0100 +++ b/src/Pure/IsaMakefile Tue Feb 01 18:01:57 2005 +0100 @@ -23,36 +23,42 @@ Pure: $(OUT)/Pure -$(OUT)/Pure: General/ROOT.ML General/buffer.ML General/file.ML \ - General/graph.ML General/heap.ML General/history.ML \ - General/lazy_scan.ML General/lazy_seq.ML General/name_space.ML \ - General/object.ML General/output.ML General/path.ML General/position.ML \ - General/pretty.ML General/scan.ML General/seq.ML General/source.ML \ - General/susp.ML General/symbol.ML General/table.ML General/url.ML \ - General/xml.ML Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML \ - Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML Isar/constdefs.ML \ - Isar/context_rules.ML Isar/induct_attrib.ML Isar/isar.ML \ - Isar/isar_cmd.ML Isar/isar_output.ML Isar/isar_syn.ML Isar/isar_thy.ML \ - Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/object_logic.ML \ - Isar/obtain.ML Isar/outer_lex.ML Isar/outer_parse.ML \ - Isar/outer_syntax.ML Isar/proof.ML Isar/proof_context.ML \ - Isar/proof_data.ML Isar/proof_history.ML Isar/rule_cases.ML \ - Isar/session.ML Isar/skip_proof.ML Isar/thy_header.ML Isar/toplevel.ML \ - ML-Systems/cpu-timer-basis.ML ML-Systems/cpu-timer-gc.ML \ - ML-Systems/polyml.ML ML-Systems/polyml-time-limit.ML \ - ML-Systems/smlnj-basis-compat.ML ML-Systems/smlnj-compiler.ML \ - ML-Systems/smlnj-pp-new.ML ML-Systems/smlnj-pp-old.ML \ - ML-Systems/smlnj.ML Proof/ROOT.ML Proof/extraction.ML \ - Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML Proof/proofchecker.ML \ - Proof/reconstruct.ML ROOT.ML Syntax/ROOT.ML Syntax/ast.ML Syntax/lexicon.ML \ - Syntax/mixfix.ML Syntax/parser.ML Syntax/printer.ML Syntax/syn_ext.ML \ - Syntax/syn_trans.ML Syntax/syntax.ML Syntax/token_trans.ML Syntax/type_ext.ML \ - Thy/ROOT.ML Thy/html.ML Thy/latex.ML Thy/present.ML Thy/thm_database.ML \ - Thy/thm_deps.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML \ - Thy/thy_scan.ML Thy/thy_syn.ML axclass.ML codegen.ML context.ML display.ML \ - drule.ML envir.ML fact_index.ML goals.ML install_pp.ML library.ML logic.ML \ - meta_simplifier.ML net.ML pattern.ML proof_general.ML proofterm.ML pure.ML \ - pure_thy.ML search.ML sign.ML sorts.ML tactic.ML tctical.ML term.ML theory.ML \ +$(OUT)/Pure: General/ROOT.ML General/buffer.ML General/file.ML \ + General/graph.ML General/heap.ML General/history.ML \ + General/lazy_scan.ML General/lazy_seq.ML General/name_space.ML \ + General/object.ML General/output.ML General/path.ML General/position.ML\ + General/pretty.ML General/scan.ML General/seq.ML General/source.ML \ + General/susp.ML General/symbol.ML General/table.ML General/url.ML \ + General/xml.ML\ + IsaPlanner/isaplib.ML IsaPlanner/term_lib.ML\ + IsaPlanner/isa_fterm.ML IsaPlanner/upterm_lib.ML\ + IsaPlanner/focus_term_lib.ML\ + IsaPlanner/rw_tools.ML IsaPlanner/rw_inst.ML\ + IsaPlanner/isand.ML\ + Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML \ + Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML Isar/constdefs.ML\ + Isar/context_rules.ML Isar/induct_attrib.ML Isar/isar.ML \ + Isar/isar_cmd.ML Isar/isar_output.ML Isar/isar_syn.ML Isar/isar_thy.ML\ + Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/object_logic.ML \ + Isar/obtain.ML Isar/outer_lex.ML Isar/outer_parse.ML \ + Isar/outer_syntax.ML Isar/proof.ML Isar/proof_context.ML \ + Isar/proof_data.ML Isar/proof_history.ML Isar/rule_cases.ML \ + Isar/session.ML Isar/skip_proof.ML Isar/thy_header.ML Isar/toplevel.ML\ + ML-Systems/cpu-timer-basis.ML ML-Systems/cpu-timer-gc.ML \ + ML-Systems/polyml.ML ML-Systems/polyml-time-limit.ML \ + ML-Systems/smlnj-basis-compat.ML ML-Systems/smlnj-compiler.ML \ + ML-Systems/smlnj-pp-new.ML ML-Systems/smlnj-pp-old.ML \ + ML-Systems/smlnj.ML Proof/ROOT.ML Proof/extraction.ML \ + Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML Proof/proofchecker.ML\ + Proof/reconstruct.ML ROOT.ML Syntax/ROOT.ML Syntax/ast.ML Syntax/lexicon.ML\ + Syntax/mixfix.ML Syntax/parser.ML Syntax/printer.ML Syntax/syn_ext.ML \ + Syntax/syn_trans.ML Syntax/syntax.ML Syntax/token_trans.ML Syntax/type_ext.ML\ + Thy/ROOT.ML Thy/html.ML Thy/latex.ML Thy/present.ML Thy/thm_database.ML\ + Thy/thm_deps.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML \ + Thy/thy_scan.ML Thy/thy_syn.ML axclass.ML codegen.ML context.ML display.ML\ + drule.ML envir.ML fact_index.ML goals.ML install_pp.ML library.ML logic.ML\ + meta_simplifier.ML net.ML pattern.ML proof_general.ML proofterm.ML pure.ML\ + pure_thy.ML search.ML sign.ML sorts.ML tactic.ML tctical.ML term.ML theory.ML\ theory_data.ML thm.ML type.ML type_infer.ML unify.ML @./mk diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/ROOT.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,15 @@ +(* ID: $Id$ + Author: Lucas Dixon, University of Edinburgh + lucasd@dai.ed.ac.uk + +The IsaPlanner subsystem. +*) + +use "isand.ML"; +use "isaplib.ML"; +use "term_lib.ML"; +use "upterm_lib.ML"; +use "focus_term_lib.ML"; +use "rw_tools.ML"; +use "rw_inst.ML"; +use "isa_fterm.ML"; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/focus_term_lib.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/focus_term_lib.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,379 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: libs/focus_term_lib.ML + Author: Lucas Dixon, University of Edinburgh + lucasd@dai.ed.ac.uk + Date: 16 April 2003 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Generic Foucs terms (like Zippers), which allows modification and + manipulation of term keeping track of how we got to the position + in the term. We provide a signature for terms, ie any kind of + lamda calculus with abs and application, and some notion of types + and naming of abstracted vars. + + FIXME: at some point in the future make a library to work simply + with polymorphic upterms - that way if I want to use them without + the focus part, then I don't need to include all the focus stuff. + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +(* to endoce and decode general terms into the type needed by focus +term, this is an easy thing to write, but needs to be done for each +type that is going to be used with focus terms. *) + +signature F_ENCODE_TERM_SIG = +sig + + type term (* type of term to be encoded into TermT for focus term manip *) + type TypeT (* type annotation for abstractions *) + type LeafT (* type for other leaf types in term sturcture *) + + (* the type to be encoded into *) + datatype TermT = $ of TermT * TermT + | Abs of string * TypeT * TermT + | lf of LeafT + + (* the encode and decode functions *) + val fakebounds : string * TypeT -> term -> term + val encode : term -> TermT + val decode : TermT -> term + +end; + + + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +signature FOCUS_TERM_SIG = +sig + structure MinTermS : F_ENCODE_TERM_SIG + + exception focus_term_exp of string; + exception out_of_term_exception of string; + + type Term = MinTermS.TermT; + type Type = MinTermS.TypeT; + + type UpTerm = (Term,Type) UpTermLib.T; + +(* datatype + UpTerm = + abs of string * Type * UpTerm + | appl of Term * UpTerm + | appr of Term * UpTerm + | root *) + + datatype FcTerm = focus of Term * UpTerm + + (* translating *) + val fcterm_of_term : MinTermS.term -> FcTerm + val term_of_fcterm : FcTerm -> MinTermS.term + + (* editing/constrution *) + val enc_appl : (MinTermS.term * UpTerm) -> UpTerm + val enc_appr : (MinTermS.term * UpTerm) -> UpTerm + + (* upterm creatioin *) + val upterm_of : FcTerm -> UpTerm + val add_upterm : UpTerm -> FcTerm -> FcTerm + val mk_term_of_upterm : (MinTermS.term * UpTerm) -> MinTermS.term + val mk_termf_of_upterm : UpTerm -> + (((string * Type) list) * + (MinTermS.term -> MinTermS.term)) + val pure_mk_termf_of_upterm : (MinTermS.term, Type) UpTermLib.T -> + (((string * Type) list) * + (MinTermS.term -> MinTermS.term)) + + (* focusing, throws exceptions *) + val focus_bot_left_leaf : FcTerm -> FcTerm + val focus_left : FcTerm -> FcTerm + val focus_right : FcTerm -> FcTerm + val focus_abs : FcTerm -> FcTerm + val focus_fake_abs : FcTerm -> FcTerm + val focus_to_top : FcTerm -> FcTerm + val focus_up : FcTerm -> FcTerm + val focus_up_right : FcTerm -> FcTerm + val focus_up_right1 : FcTerm -> FcTerm + + (* optional focus changes *) + val focus_up_abs : FcTerm -> FcTerm option + val focus_up_appr : FcTerm -> FcTerm option + val focus_up_appl : FcTerm -> FcTerm option + val focus_up_abs_or_appr : FcTerm -> FcTerm option + + val tyenv_of_focus : FcTerm -> (string * Type) list + val tyenv_of_focus' : FcTerm -> Type list + + (* set/get the focus term *) + val set_focus_of_fcterm : FcTerm -> MinTermS.term -> FcTerm + val focus_of_fcterm : FcTerm -> MinTermS.term + + (* leaf navigation *) + val leaf_seq_of_fcterm : FcTerm -> FcTerm Seq.seq + val next_leaf_fcterm : FcTerm -> FcTerm + val next_leaf_of_fcterm_seq : FcTerm -> FcTerm Seq.seq + + (* analysis *) + val upsize_of_fcterm : FcTerm -> int + val pretty_fcterm : FcTerm -> Pretty.T +end; + + + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* + + NOTE!!! I think this can be done in a purely functional way with a + pair of a term (the focus) and a function, that when applied + unfocuses one level... maybe think about this as it would be a very + nice generic consrtuction, without the need for encoding/decoding + strutcures. + +*) +functor FocusTermFUN(structure EncodeTerm : F_ENCODE_TERM_SIG) + : FOCUS_TERM_SIG = +struct + +structure MinTermS = EncodeTerm; + +local open MinTermS open UpTermLib in + + type Term = MinTermS.TermT; + type Type = MinTermS.TypeT; + + exception focus_term_exp of string; + exception out_of_term_exception of string; + + infix 9 $ + + (* datatype to hold a term tree and the path to where you are in the term *) +(* datatype UpTerm = root + | appl of (Term * UpTerm) + | appr of (Term * UpTerm) + | abs of (string * Type * UpTerm); *) + + type UpTerm = (Term,Type) UpTermLib.T; + + fun enc_appl (t,u) = appl((MinTermS.encode t),u); + fun enc_appr (t,u) = appr((MinTermS.encode t),u); + + datatype FcTerm = focus of (Term * UpTerm); + + (* the the term of the upterm *) + fun mk_term_of_upt_aux (t, root) = MinTermS.decode t + | mk_term_of_upt_aux (t, appl (l,m)) = mk_term_of_upt_aux(l$t, m) + | mk_term_of_upt_aux (t, appr (r,m)) = mk_term_of_upt_aux(t$r, m) + | mk_term_of_upt_aux (t, abs (s,ty,m)) = mk_term_of_upt_aux(Abs(s,ty,t),m); + + fun mk_term_of_upterm (t, ut) = mk_term_of_upt_aux (MinTermS.encode t, ut); + + (* a function version of the above, given an upterm it returns: + a function on that given a term puts it in the context of the upterm, + and a list of bound variables that are in the upterm, added as + we go up - so in reverse order from that typiclaly used by top-down + parsing of terms. *) + fun mk_termf_of_upt_aux (f, bs, root) = + (bs, fn t => MinTermS.decode (f t)) + | mk_termf_of_upt_aux (f, bs, appl (l,m)) = + mk_termf_of_upt_aux (fn t => l $ (f t), bs, m) + | mk_termf_of_upt_aux (f, bs, appr (r,m)) = + mk_termf_of_upt_aux (fn t => (f t) $ r, bs, m) + | mk_termf_of_upt_aux (f, bs, abs (s,ty,m)) = + mk_termf_of_upt_aux (fn t => Abs(s,ty,(f t)), (s,ty)::bs, m); + + fun mk_termf_of_upterm ut = mk_termf_of_upt_aux (MinTermS.encode, [], ut); + + + fun upterm_of (focus(t, u)) = u; + + (* put a new upterm at the start of our current term *) + fun add_upterm u2 (focus(t, u)) = focus(t, UpTermLib.apply u2 u); + + + (* As above but work on the pure originial upterm type *) + fun pure_mk_termf_of_upterm ut = + mk_termf_of_upt_aux + (encode, [], (map_to_upterm_parts (encode, I) ut)); + + fun fcterm_of_term t = focus(encode t, root); + + fun term_of_fcterm (focus (t, m)) = mk_term_of_upt_aux (t, m); + +(* fun term_of_fcterm (focus (t, root)) = mk_term_of_upt_aux (t, root) + | term_of_fcterm _ = raise focus_term_exp "term_of_fcterm: something bad has happened here."; *) + + fun focus_of_fcterm (focus(t, _)) = MinTermS.decode t; + + (* replace the focus term with a new term... *) + fun set_focus_of_fcterm (focus(_, m)) nt = focus(MinTermS.encode nt, m); + + fun focus_left (focus(a $ b, m)) = focus(a, appr(b, m)) + | focus_left (focus(Abs(s,ty,t), m)) = focus_left (focus(t, abs(s,ty,m))) + | focus_left (focus(l, m)) = raise out_of_term_exception "focus_left"; + +(* Note: ":" is standard for faked bound vars... see: EQStepTacTermTools *) + fun focus_fake_abs (focus(Abs(s,ty,t), m)) = + let val t' = MinTermS.encode (fakebounds (s,ty) (MinTermS.decode t)) + in focus(t', abs(s,ty,m)) end + | focus_fake_abs (focus(l, m)) = raise out_of_term_exception "focus_fake_abs"; + + fun focus_abs (focus(Abs(s,ty,t), m)) = focus(t, abs(s,ty,m)) + | focus_abs (focus(l, m)) = raise out_of_term_exception "focus_abs"; + + + fun focus_right (focus(a $ b, m)) = focus(b, appl(a, m)) + | focus_right (focus(Abs(s,ty,t), m)) = focus_right (focus(t, abs(s,ty,m))) + | focus_right (focus(l, m)) = raise out_of_term_exception "focus_right"; + + fun focus_up (focus(t, appl(l,m))) = focus(l$t, m) + | focus_up (focus(t, appr(r,m))) = focus(t$r, m) + | focus_up (focus(t, abs(s,ty,m))) = focus_up (focus(Abs(s,ty,t), m)) + | focus_up (focus(t, root)) = raise out_of_term_exception "focus_up"; + + fun focus_to_top t = + focus_to_top (focus_up t) handle out_of_term_exception _ => t; + + (* focus up until you can move right, and then do so *) + fun focus_up_right (focus(t, appl(l,m))) = + focus_up_right (focus(l$t, m)) + | focus_up_right (focus(t, appr(r,m))) = + focus(r, appl(t,m)) + | focus_up_right (focus(t, abs(s,ty,m))) = + focus_up_right (focus(Abs(s,ty,t), m)) + | focus_up_right (focus(t, root)) = + raise out_of_term_exception "focus_up_right"; + + (* as above, but do not move up over a left application *) + fun focus_up_right1 (focus(t, appl(l,m))) = + raise out_of_term_exception "focus_up_right1" + | focus_up_right1 (focus(t, appr(r,m))) = + focus(r, appl(t,m)) + | focus_up_right1 (focus(t, abs(s,ty,m))) = + focus_up_right (focus(Abs(s,ty,t), m)) + | focus_up_right1 (focus(t, root)) = + raise out_of_term_exception "focus_up_right1"; + + (* move focus to the bottom left *) + fun focus_bot_left_leaf (ft as focus(t, _)) = + focus_bot_left_leaf (focus_left ft) + handle out_of_term_exception _=> ft; + + (* focus tools for working directly with the focus representation *) + fun focus_up_appr (focus(t, appl(l,m))) = None + | focus_up_appr (focus(t, appr(r,m))) = Some (focus(t$r, m)) + | focus_up_appr (focus(t, abs(s,ty,m))) = None + | focus_up_appr (focus(t, root)) = None; + + fun focus_up_appl (focus(t, appl(l,m))) = Some (focus(l$t, m)) + | focus_up_appl (focus(t, appr(r,m))) = None + | focus_up_appl (focus(t, abs(s,ty,m))) = None + | focus_up_appl (focus(t, root)) = None; + + fun focus_up_abs (focus(t, appl(l,m))) = None + | focus_up_abs (focus(t, appr(r,m))) = None + | focus_up_abs (focus(t, abs(s,ty,m))) = + Some (focus_up (focus(Abs(s,ty,t), m))) + | focus_up_abs (focus(t, root)) = None; + + fun focus_up_abs_or_appr (focus(t, appl(l,m))) = None + | focus_up_abs_or_appr (focus(t, appr(r,m))) = Some (focus(t$r, m)) + | focus_up_abs_or_appr (focus(t, abs(s,ty,m))) = + Some (focus_up (focus(Abs(s,ty,t), m))) + | focus_up_abs_or_appr (focus(t, root)) = None; + + + fun tyenv_of_focus (focus(t, u)) = tyenv_of_upterm u; + fun tyenv_of_focus' (focus(t, u)) = tyenv_of_upterm' u; + + (* return the Term.type of the focus term, computes bound vars type, + does a quick check only. *) +(* fun fastype_of_focus (focus(t, m)) = + let + fun boundtypes_of_upterm (abs(s,ty,m)) = + ty::(boundtypes_of_upterm m) + | boundtypes_of_upterm (appl(t,m)) = + boundtypes_of_upterm m + | boundtypes_of_upterm (appr(t,m)) = + boundtypes_of_upterm m + | boundtypes_of_upterm (root) = [] + in + fastype_of1 ((boundtypes_of_upterm m), t) + end; *) + + (* gets the next left hand side term, or throws an exception *) + fun next_leaf_fcterm ft = focus_bot_left_leaf (focus_up_right ft); + + fun next_leaf_of_fcterm_seq_aux t () = + let val nextt = next_leaf_fcterm t + in + Some (nextt, Seq.make (next_leaf_of_fcterm_seq_aux nextt)) + end handle out_of_term_exception _ => None; + + (* sequence of next upterms from start upterm... + ie a sequence of all leafs after this one*) + fun next_leaf_of_fcterm_seq in_t = + Seq.make (next_leaf_of_fcterm_seq_aux in_t); + + (* returns a sequence of all leaf up terms from term, ie DFS on + leafs of term, ie uses time O(n), n = sizeof term. *) + fun leaf_seq_of_fcterm in_t = + let + val botleft = (focus_bot_left_leaf in_t) + in + Seq.cons (botleft, (Seq.make (next_leaf_of_fcterm_seq_aux botleft)) ) + end; + + + fun upsize_of_fcterm (focus(t, ut)) = upsize_of_upterm ut; + + fun pretty_fcterm ft = Pretty.str "no yet implemented"; + +end; (* local *) + +end; (* functor *) + + + + + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* focus term encoding sturcture for isabelle terms *) +structure EncodeIsaFTerm : F_ENCODE_TERM_SIG = +struct + infix 9 $; + + type term = Term.term; + + type TypeT = Term.typ; + + datatype LeafT = lConst of string * Term.typ + | lVar of ((string * int) * Term.typ) + | lFree of (string * Term.typ) + | lBound of int; + + datatype TermT = op $ of TermT * TermT + | Abs of string * TypeT * TermT + | lf of LeafT; + + fun fakebounds (s, ty) t = subst_bound (Free(":" ^ s, ty), t); + + fun encode (Term.$(a,b)) = (encode a) $ (encode b) + | encode (Term.Abs(s,ty,t)) = Abs(s,ty,(encode t)) + | encode (Term.Const(s,ty)) = lf(lConst(s,ty)) + | encode (Term.Var((s,i),ty)) = lf(lVar((s,i),ty)) + | encode (Term.Free(s,ty)) = lf(lFree(s,ty)) + | encode (Term.Bound i) = lf(lBound i); + + fun decode (a $ b) = Term.$(decode a, decode b) + | decode (Abs(s,ty,t)) = (Term.Abs(s,ty,decode t)) + | decode (lf(lConst(s,ty))) = Term.Const(s,ty) + | decode (lf(lVar((s,i),ty))) = Term.Var((s,i),ty) + | decode (lf(lFree(s,ty))) = (Term.Free(s,ty)) + | decode (lf(lBound i)) = (Term.Bound i); + +end; + diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/isa_fterm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/isa_fterm.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,368 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: libs/isa_fterm.ML + Author: Lucas Dixon, University of Edinburgh + lucasd@dai.ed.ac.uk + Date: 16 April 2003 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Generic Foucs terms (like Zippers) instantiation for Isabelle terms. + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +signature ISA_ENCODE_TERM = + F_ENCODE_TERM_SIG + where type term = Term.term type TypeT = Term.typ; + + +(* signature BASIC_ISA_FTERM = +FOCUS_TERM_SIG where type Term = ISA_ENCODE_TERM.Term *) + +(* +signature ISA_FTERM = +sig + include F_ENCODE_TERM_SIG + + val clean_match_ft : + Type.tsig -> + FcTerm -> + Term.term -> + ( + ((Term.indexname * Term.typ) list * (Term.indexname * Term.term) list) + * (string * Type) list * Term.term) Library.option + val clean_unify_ft : + Sign.sg -> + int -> + FcTerm -> + Term.term -> + ( + ((Term.indexname * Term.typ) list * (Term.indexname * Term.term) list) + * (string * Type) list * Term.term) Seq.seq + val fakefree_badbounds : + (string * Term.typ) list -> int -> Term.term -> Term.term + val find_fcterm_matches : + ((FcTerm -> 'a) -> FcTerm -> 'b) -> + (FcTerm -> 'a) -> FcTerm -> 'b + val find_sg_concl_matches : + ((FcTerm -> 'a) -> FcTerm -> 'b) -> + (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b + val find_sg_concl_thm_matches : + ((FcTerm -> 'a) -> FcTerm -> 'b) -> + (FcTerm -> 'a) -> int -> Thm.thm -> 'b + val find_sg_matches : + ((FcTerm -> 'a) -> FcTerm -> 'b) -> + (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b + val find_sg_thm_matches : + ((FcTerm -> 'a) -> FcTerm -> 'b) -> + (FcTerm -> 'a) -> int -> Thm.thm -> 'b + val focus_to_concl : FcTerm -> FcTerm + val focus_to_concl_of_term : EncodeIsaFTerm.term -> FcTerm + val focus_to_subgoal : + int -> FcTerm -> FcTerm + val focus_to_subgoal_of_term : + int -> EncodeIsaFTerm.term -> FcTerm + val focus_to_term_goal_prem : + int * int -> EncodeIsaFTerm.term -> FcTerm + val focuseq_to_subgoals : + FcTerm -> FcTerm Seq.seq + exception isa_focus_term_exp of string + val mk_foo_match : + (Term.term -> Term.term) -> + ('a * Term.typ) list -> Term.term -> Term.term + val prepmatch : + FcTerm -> + Term.term * ((string * Type) list * Term.term) + val search_bl_ru_f : + (FcTerm -> 'a Seq.seq) -> + FcTerm -> 'a Seq.seq + val search_bl_ur_f : + (FcTerm -> 'a Seq.seq) -> + FcTerm -> 'a Seq.seq + val valid_match_start : FcTerm -> bool + +end *) + +structure BasicIsaFTerm = +FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm); + +(* HOL Dependent *) +structure IsaFTerm = +struct + +(* include BasicIsaFTerm *) +(* open FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm ); *) + +open BasicIsaFTerm; + + +(* Some general search within a focus term... *) + +(* Note: only upterms with a free or constant are going to yeald a +match, thus if we get anything else (bound or var) skip it! This is +important if we switch to a unification net! in particular to avoid +vars. *) + +fun valid_match_start ft = + (case TermLib.bot_left_leaf_of (focus_of_fcterm ft) of + Const _ => true + | Free _ => true + | Abs _ => true + | _ => false); + +(* match starting at the bottom left, moving up to top of the term, + then moving right to the next leaf and up again etc *) +(* FIXME: make properly lazy! *) +fun search_nonvar_bl_ur_f f ft = + let + val fts = + Seq.filter valid_match_start + (leaf_seq_of_fcterm ft) + + fun mk_match_up_seq ft = + case focus_up_abs_or_appr ft of + Some ft' => Seq.append(f ft, mk_match_up_seq ft') + | None => f ft + in + Seq.flat (Seq.map mk_match_up_seq fts) + end; + +fun search_bl_ur_f f ft = + let + val fts = (leaf_seq_of_fcterm ft) + + fun mk_match_up_seq ft = + case focus_up_abs_or_appr ft of + Some ft' => Seq.append(f ft, mk_match_up_seq ft') + | None => f ft + in + Seq.flat (Seq.map mk_match_up_seq fts) + end; + + +(* FIXME: make properly lazy! *) +(* FIXME: make faking of bound vars local - for speeeeed *) +fun search_nonvar_bl_ru_f f ft = + let + fun mauxtop ft = + if (valid_match_start ft) + then maux ft else Seq.empty + and maux ft = + let val t' = (focus_of_fcterm ft) + (* val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *) + in + (case t' of + (_ $ _) => Seq.append (maux (focus_left ft), + Seq.append(mauxtop (focus_right ft), + f ft)) + | (Abs _) => Seq.append (maux (focus_abs ft), + f ft) + | leaf => f ft) end + in + mauxtop ft + end; + + +fun search_bl_ru_f f ft = + let + fun maux ft = + let val t' = (focus_of_fcterm ft) + (* val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *) + in + (case t' of + (_ $ _) => Seq.append (maux (focus_left ft), + Seq.append(maux (focus_right ft), + f ft)) + | (Abs _) => Seq.append (maux (focus_abs ft), + f ft) + | leaf => f ft) end + in maux ft end; + + + +exception isa_focus_term_exp of string; + + +fun focus_to_dest_impl ft = + let val (lhs, rhs) = + Logic.dest_implies (focus_of_fcterm ft) + in (focus_left ft, focus_right ft) end + handle Term.TERM _ => + raise isa_focus_term_exp + "focus_to_dest_impl: applied to non implication"; + + +(* move focus to the conlusion *) +fun focus_to_concl ft = + let val (lhs, rhs) = + Logic.dest_implies (focus_of_fcterm ft) + in focus_to_concl (focus_right ft) end + handle Term.TERM _ => ft; + +val focus_to_concl_of_term = focus_to_concl o fcterm_of_term; + + + +(* give back sequence of focuses at different subgoals *) +(* FIXME: make truly lazy *) +fun focuseq_to_subgoals ft = + if (Logic.is_implies (focus_of_fcterm ft)) then + Seq.cons (focus_right (focus_left ft), focuseq_to_subgoals (focus_right ft)) + else + Seq.empty; + +(* move focus to a specific subgoal, 0 is first *) +fun focus_to_subgoal j ft = + let fun focus_to_subgoal' (ft, 0) = + let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) + in ft |> focus_left |> focus_right end + | focus_to_subgoal' (ft, i) = + let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) + in focus_to_subgoal' (focus_right ft, i - 1) end + in focus_to_subgoal' (ft, j - 1) end + handle Term.TERM _ => + raise isa_focus_term_exp + ("focus_to_subgoal: No such subgoal: " ^ + (string_of_int j)); + +fun focus_to_subgoal_of_term i t = + focus_to_subgoal i (fcterm_of_term t) + +(* move focus to a specific premise *) +(* fun focus_past_params i ft = + (focus_to_subgoal (focus_right ft, i)) + handle isa_focus_term_exp _ => + raise isa_focus_term_exp + ("focus_to_prmise: No such premise: " ^ (string_of_int i)); *) + +fun focus_to_term_goal_prem (premid,gaolid) t = + focus_to_subgoal premid (focus_to_subgoal_of_term gaolid t); + + + + +(* T is outer bound vars, n is number of locally bound vars *) +fun fakefree_badbounds T n (a $ b) = + (fakefree_badbounds T n a) $ (fakefree_badbounds T n b) + | fakefree_badbounds T n (Abs(s,ty,t)) = + Abs(s,ty, fakefree_badbounds T (n + 1) t) + | fakefree_badbounds T n (b as Bound i) = + let fun mkfake_bound j [] = + raise ERROR_MESSAGE "fakefree_badbounds: bound is outside of the known types!" + | mkfake_bound 0 ((s,ty)::Ts) = + Free (RWTools.mk_fake_bound_name s,ty) + | mkfake_bound j (d::Ts) = mkfake_bound (j - 1) Ts + in if n <= i then mkfake_bound (i - n) T else b end + | fakefree_badbounds T n t = t; + + +(* note: outerterm is the taget with the match replaced by a bound + variable : ie: "P lhs" beocmes "%x. P x" + insts is the types of instantiations of vars in lhs + and typinsts is the type instantiations of types in the lhs + Note: Final rule is the rule lifted into the ontext of the + taget thm. *) +fun mk_foo_match mkuptermfunc Ts t = + let + val ty = Term.type_of t + val bigtype = (rev (map snd Ts)) ---> ty + fun mk_foo 0 t = t + | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) + val num_of_bnds = (length Ts) + (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) + val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) + in Abs("fooabs", bigtype, mkuptermfunc foo_term) end; + +(* before matching we need to fake the bound vars that missing an +abstraction. In this function we additionally construct the +abstraction environment, and an outer context term (with the focus +abstracted out) for use in rewriting with RWInst.rw *) +fun prepmatch ft = + let + val t = focus_of_fcterm ft + val Ts = tyenv_of_focus ft + val t' = fakefree_badbounds Ts 0 t + fun mktermf t = + term_of_fcterm (set_focus_of_fcterm ft t) + val absterm = mk_foo_match mktermf Ts t' + in + (t', (Ts, absterm)) + end; + +(* matching and unification for a focus term's focus *) +fun clean_match_ft tsig pat ft = + let val (t, (Ts,absterm)) = prepmatch ft in + case TermLib.clean_match tsig (pat, t) of + None => None + | Some insts => Some (insts, Ts, absterm) end; +(* ix = max var index *) +fun clean_unify_ft sgn ix pat ft = + let val (t, (Ts,absterm)) = prepmatch ft in + Seq.map (fn insts => (insts, Ts, absterm)) + (TermLib.clean_unify sgn ix (t, pat)) end; + + +(* THINK: ? do we not need to incremement bound indices? *) +(* THINK: it should be the search which localisaes the search to the +current focus, not this hack in the matcher... ? *) +(* find matches below this particular focus term *) +(* The search function is to find a match within a term... *) +(* the matcher is something that is applied to each node chosen by the +searchf and the results are flattened to form a lazy list. *) +fun find_fcterm_matches searchf matcher ft = + let + val ftupterm = upterm_of ft + val focusft = focus_of_fcterm ft + val add_uptermf = add_upterm ftupterm + in + searchf + (fn ft' => matcher (add_uptermf ft')) + (fcterm_of_term focusft) + end; + +(* FIXME: move argument orders for efficiency... +i.e. wenzel style val foofunc = x o y; +*) + +(* find the matches inside subgoal i of th *) +fun find_sg_matches searchf matcher i t = + let val subgoal_fcterm = focus_to_subgoal_of_term i t + in find_fcterm_matches searchf matcher subgoal_fcterm end; + +(* find the matches inside subgoal i of th *) +fun find_sg_thm_matches searchf matcher i th = + find_sg_matches searchf matcher i (Thm.prop_of th); + + +(* find the matches inside subgoal i's conclusion of th *) +fun find_sg_concl_matches searchf matcher i t = + let + val subgoal_fcterm = + focus_to_concl (focus_to_subgoal_of_term i t) + in + find_fcterm_matches searchf matcher subgoal_fcterm + end; + +(* find the matches inside subgoal i's conclusion of th *) +fun find_sg_concl_thm_matches searchf matcher i th = + find_sg_concl_matches searchf matcher i (Thm.prop_of th); + +end; + +(* +test... + +f_encode_isatermS.encode (read "P a"); +isafocustermS.fcterm_of_term (read "f a"); +isafocustermS.term_of_fcterm it; + +Goal "P b ==> P (suc b)"; + +TermLib.string_of_term ((focus_of_fcterm o focus_to_subgoal_of_term 1 o prop_of) (topthm())); + +TermLib.string_of_term ((focus_of_fcterm o focus_to_concl o focus_to_subgoal_of_term 1 o prop_of) (topthm())); + +TermLib.string_of_term ((focus_of_fcterm o focus_to_term_goal_prem (1,1) o prop_of) (topthm())); + +*) diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/isand.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/isand.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,486 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/isand.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Date: 6 Aug 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Natural Deduction tools + + For working with Isbaelle theorem in a natural detuction style, + ie, not having to deal with meta level quantified varaibles, + instead, we work with newly introduced frees, and hide the + "all"'s, exporting results from theorems proved with the frees, to + solve the all cases of the previous goal. + + Note: A nice idea: allow esxporting to solve any subgoal, thus + allowing the interleaving of proof, or provide a structure for the + ordering of proof, thus allowing proof attempts in parrelle, but + recording the order to apply things in. + + debugging tools: + + fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); + fun asm_read s = + (assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); + + THINK: are we really ok with our varify name w.r.t the prop - do + we alos need to avoid names in the hidden hyps? +*) + +signature ISA_ND = +sig + datatype export = export of + {gth: Thm.thm, (* initial goal theorem *) + sgid: int, (* subgoal id which has been fixed etc *) + fixes: Thm.cterm list, (* frees *) + assumes: Thm.cterm list} (* assumptions *) + + val fixes_of_exp : export -> Thm.cterm list + + val export_back : export -> Thm.thm -> Thm.thm Seq.seq + val export_solution : export -> Thm.thm -> Thm.thm + val export_solutions : export list * Thm.thm -> Thm.thm + + val allify_conditions : + (Term.term -> Thm.cterm) -> + (string * Term.typ) list -> Thm.thm -> Thm.thm * Thm.cterm list + val allify_conditions' : + (string * Term.typ) list -> Thm.thm -> Thm.thm * Thm.cterm list + + val assume_prems : + int -> Thm.thm -> Thm.thm list * Thm.thm * Thm.cterm list + + val fix_alls_term : int -> Term.term -> Term.term * Term.term list + val fix_alls_cterm : int -> Thm.thm -> Thm.cterm * Thm.cterm list + val fix_alls' : int -> Thm.thm -> Thm.thm * Thm.cterm list + val fix_alls : int -> Thm.thm -> Thm.thm * export + + val fix_vars_to_frees : Thm.thm -> (Thm.cterm * Thm.cterm) list * Thm.thm + val schemify_frees_to_vars : Thm.cterm list -> Thm.thm -> Thm.thm + + val fixed_subgoal_thms : Thm.thm -> Thm.thm list * (Thm.thm list -> Thm.thm) + val fixes_and_assumes : int -> Thm.thm -> Thm.thm list * Thm.thm * export + + val hide_other_goals : Thm.thm -> Thm.thm * Thm.cterm list + val hide_prems : Thm.thm -> Thm.thm * Thm.cterm list + + (* abstracts cterms (vars) to locally meta-all bounds *) + val prepare_goal_export : string list * Thm.cterm list -> Thm.thm + -> int * Thm.thm + val solve_with : Thm.thm -> Thm.thm -> Thm.thm Seq.seq + val subgoal_thms : Thm.thm -> Thm.thm list * (Thm.thm list -> Thm.thm) +end + + +structure IsaND : ISA_ND = +struct + +(* Solve *some* subgoal of "th" directly by "sol" *) +(* Note: this is probably what Markus ment to do upon export of a +"show" but maybe he used RS/rtac instead, which would wrongly lead to +failing if there are premices to the shown goal. *) +fun solve_with sol th = + let fun solvei 0 = Seq.empty + | solvei i = + Seq.append (bicompose false (false,sol,0) i th, + solvei (i - 1)) + in + solvei (Thm.nprems_of th) + end; + + +(* Given ctertmify function, (string,type) pairs capturing the free +vars that need to be allified in the assumption, and a theorem with +assumptions possibly containing the free vars, then we give back the +assumptions allified as hidden hyps. *) +(* +Given: vs +th: A vs ==> B vs +Results in: "B vs" [!!x. A x] +*) +fun allify_conditions ctermify Ts th = + let + val premts = Thm.prems_of th; + + fun allify_prem_var (vt as (n,ty),t) = + (Term.all ty) $ (Abs(n,ty,Term.abstract_over (Free vt, t))) + + fun allify_prem Ts p = foldr allify_prem_var (Ts, p) + + val cTs = map (ctermify o Free) Ts + val cterm_asms = map (ctermify o allify_prem Ts) premts + val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms + in + (foldl (fn (x,y) => y COMP x) (th, allifyied_asm_thms), cterm_asms) + end; + +fun allify_conditions' Ts th = + allify_conditions (Thm.cterm_of (Thm.sign_of_thm th)) Ts th; + + + +(* change schematic vars to fresh free vars *) +fun fix_vars_to_frees th = + let + val ctermify = Thm.cterm_of (Thm.sign_of_thm th) + val prop = Thm.prop_of th + val vars = map Term.dest_Var (Term.term_vars prop) + + val names = Term.add_term_names (prop, []) + + val (insts,names2) = + foldl (fn ((insts,names),v as ((n,i),ty)) => + let val n2 = Term.variant names n in + ((ctermify (Var v), ctermify (Free(n2,ty))) :: insts, + n2 :: names) + end) + (([],names), vars) + in (insts, Thm.instantiate ([], insts) th) end; + +(* make free vars into schematic vars with index zero *) + fun schemify_frees_to_vars frees = + apply (map (K (Drule.forall_elim_var 0)) frees) + o Drule.forall_intr_list frees; + + +(* datatype to capture an exported result, ie a fix or assume. *) +datatype export = + export of {fixes : Thm.cterm list, (* fixed vars *) + assumes : Thm.cterm list, (* hidden hyps/assumed prems *) + sgid : int, + gth : Thm.thm}; (* subgoal/goalthm *) + +fun fixes_of_exp (export rep) = #fixes rep; + +(* export the result of the new goal thm, ie if we reduced teh +subgoal, then we get a new reduced subtgoal with the old +all-quantified variables *) +local + +(* allify puts in a meta level univ quantifier for a free variavble *) +fun allify_term (v, t) = + let val vt = #t (Thm.rep_cterm v) + val (n,ty) = Term.dest_Free vt + in (Term.all ty) $ (Abs(n,ty,Term.abstract_over (vt, t))) end; + +fun allify_for_sg_term ctermify vs t = + let val t_alls = foldr allify_term (vs,t); + val ct_alls = ctermify t_alls; + in + (ct_alls, Drule.forall_elim_list vs (Thm.assume ct_alls)) + end; +(* lookup type of a free var name from a list *) +fun lookupfree vs vn = + case Library.find_first (fn (n,ty) => n = vn) vs of + None => raise ERROR_MESSAGE ("prepare_goal_export:lookupfree: " + ^ vn ^ " does not occur in the term") + | Some x => x; +in +fun export_back (export {fixes = vs, assumes = hprems, + sgid = i, gth = gth}) newth = + let + val sgn = Thm.sign_of_thm newth; + val ctermify = Thm.cterm_of sgn; + + val sgs = prems_of newth; + val (sgallcts, sgthms) = + Library.split_list (map (allify_for_sg_term ctermify vs) sgs); + val minimal_newth = + (foldl (fn ( newth', sgthm) => + Drule.compose_single (sgthm, 1, newth')) + (newth, sgthms)); + val allified_newth = + minimal_newth + |> Drule.implies_intr_list hprems + |> Drule.forall_intr_list vs + + val newth' = Drule.implies_intr_list sgallcts allified_newth + in + bicompose false (false, newth', (length sgallcts)) i gth + end; + +(* +Given "vs" : names of free variables to abstract over, +Given cterms : premices to abstract over (P1... Pn) in terms of vs, +Given a thm of the form: +P1 vs; ...; Pn vs ==> Goal(C vs) + +Gives back: +(n, length of given cterms which have been allified + [| !! vs. P1 vs; !! vs. Pn vs |] ==> !! C vs) the allified thm +*) +(* note: C may contain further premices etc +Note that cterms is the assumed facts, ie prems of "P1" that are +reintroduced in allified form. +*) +fun prepare_goal_export (vs, cterms) th = + let + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + + val allfrees = map Term.dest_Free (Term.term_frees (Thm.prop_of th)) + val cfrees = map (ctermify o Free o lookupfree allfrees) vs + + val sgs = prems_of th; + val (sgallcts, sgthms) = + Library.split_list (map (allify_for_sg_term ctermify cfrees) sgs); + + val minimal_th = + (foldl (fn ( th', sgthm) => + Drule.compose_single (sgthm, 1, th')) + (th, sgthms)) RS Drule.rev_triv_goal; + + val allified_th = + minimal_th + |> Drule.implies_intr_list cterms + |> Drule.forall_intr_list cfrees + + val th' = Drule.implies_intr_list sgallcts allified_th + in + ((length sgallcts), th') + end; + +end; + + +(* exporting function that takes a solution to the fixed/assumed goal, +and uses this to solve the subgoal in the main theorem *) +fun export_solution (export {fixes = cfvs, assumes = hcprems, + sgid = i, gth = gth}) solth = + let + val solth' = + solth |> Drule.implies_intr_list hcprems + |> Drule.forall_intr_list cfvs + in Drule.compose_single (solth', i, gth) end; + +val export_solutions = foldr (uncurry export_solution); + + +(* fix parameters of a subgoal "i", as free variables, and create an +exporting function that will use the result of this proved goal to +show the goal in the original theorem. + +Note, an advantage of this over Isar is that it supports instantiation +of unkowns in the earlier theorem, ie we can do instantiation of meta +vars! *) +(* loosely corresponds to: +Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm +Result: + ("(As ==> SGi x') ==> (As ==> SGi x')" : thm, + expf : + ("As ==> SGi x'" : thm) -> + ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm) +*) +fun fix_alls_term i t = + let + val names = Term.add_term_names (t,[]); + val gt = Logic.get_goal t i; + val body = Term.strip_all_body gt; + val alls = rev (Term.strip_all_vars gt); + val fvs = map Free + ((Term.variantlist (map fst alls, names)) + ~~ (map snd alls)); + in ((subst_bounds (fvs,body)), fvs) end; + +fun fix_alls_cterm i th = + let + val ctermify = Thm.cterm_of (Thm.sign_of_thm th); + val (fixedbody, fvs) = fix_alls_term i (Thm.prop_of th); + val cfvs = rev (map ctermify fvs); + val ct_body = ctermify fixedbody + in + (ct_body, cfvs) + end; + +fun fix_alls' i = + (apfst Thm.trivial) o (fix_alls_cterm i); + + +(* hide other goals *) +(* note the export goal is rotated by (i - 1) and will have to be +unrotated to get backto the originial position(s) *) +fun hide_other_goals th = + let + (* tl beacuse fst sg is the goal we are interested in *) + val cprems = tl (Drule.cprems_of th) + val aprems = map Thm.assume cprems + in + (Drule.implies_elim_list (Drule.rotate_prems 1 th) aprems, + cprems) + end; + +(* a nicer version of the above that leaves only a single subgoal (the +other subgoals are hidden hyps, that the exporter suffles about) +namely the subgoal that we were trying to solve. *) +(* loosely corresponds to: +Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm +Result: + ("(As ==> SGi x') ==> SGi x'" : thm, + expf : + ("SGi x'" : thm) -> + ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm) +*) +fun fix_alls i th = + let + val (fixed_gth, fixedvars) = fix_alls' i th + val (sml_gth, othergoals) = hide_other_goals fixed_gth + in + (sml_gth, export {fixes = fixedvars, + assumes = othergoals, + sgid = i, gth = th}) + end; + + +(* assume the premises of subgoal "i", this gives back a list of +assumed theorems that are the premices of subgoal i, it also gives +back a new goal thm and an exporter, the new goalthm is as the old +one, but without the premices, and the exporter will use a proof of +the new goalthm, possibly using the assumed premices, to shoe the +orginial goal. *) + +(* Note: Dealing with meta vars, need to meta-level-all them in the +shyps, which we can later instantiate with a specific value.... ? +think about this... maybe need to introduce some new fixed vars and +then remove them again at the end... like I do with rw_inst. *) +(* loosely corresponds to: +Given "[| SG0; ... [| A0; ... An |] ==> SGi; ... SGm |] ==> G" : thm +Result: +(["A0" [A0], ... ,"An" [An]] : thm list, -- assumptions + "SGi ==> SGi" : thm, -- new goal + "SGi" ["A0" ... "An"] : thm -> -- export function + ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list) +*) +fun assume_prems i th = + let + val t = (prop_of th); + val gt = Logic.get_goal t i; + val _ = case Term.strip_all_vars gt of [] => () + | _ => raise ERROR_MESSAGE "assume_prems: goal has params" + val body = gt; + val prems = Logic.strip_imp_prems body; + val concl = Logic.strip_imp_concl body; + + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + val cprems = map ctermify prems; + val aprems = map Thm.assume cprems; + val gthi = Thm.trivial (ctermify concl); + + (* fun explortf thi = + Drule.compose (Drule.implies_intr_list cprems thi, + i, th) *) + in + (aprems, gthi, cprems) + end; + + +(* first fix the variables, then assume the assumptions *) +(* loosely corresponds to: +Given + "[| SG0; ... + !! xs. [| A0 xs; ... An xs |] ==> SGi xs; + ... SGm |] ==> G" : thm +Result: +(["A0 xs'" [A0 xs'], ... ,"An xs'" [An xs']] : thm list, -- assumptions + "SGi xs' ==> SGi xs'" : thm, -- new goal + "SGi xs'" ["A0 xs'" ... "An xs'"] : thm -> -- export function + ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list) +*) + +(* Note: the fix_alls actually pulls through all the assumptions which +means that the second export is not needed. *) +fun fixes_and_assumes i th = + let + val (fixgth, exp1) = fix_alls i th + val (assumps, goalth, _) = assume_prems 1 fixgth + in + (assumps, goalth, exp1) + end; + + +(* Fixme: allow different order of subgoals given to expf *) +(* make each subgoal into a separate thm that needs to be proved *) +(* loosely corresponds to: +Given + "[| SG0; ... SGm |] ==> G" : thm +Result: +(["SG0 ==> SG0", ... ,"SGm ==> SGm"] : thm list, -- goals + ["SG0", ..., "SGm"] : thm list -> -- export function + "G" : thm) +*) +fun subgoal_thms th = + let + val t = (prop_of th); + + val prems = Logic.strip_imp_prems t; + + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + + val aprems = map (Thm.trivial o ctermify) prems; + + fun explortf premths = + Drule.implies_elim_list th premths + in + (aprems, explortf) + end; + + +(* make all the premices of a theorem hidden, and provide an unhide +function, that will bring them back out at a later point. This is +useful if you want to get back these premices, after having used the +theorem with the premices hidden *) +(* loosely corresponds to: +Given "As ==> G" : thm +Result: ("G [As]" : thm, + "G [As]" : thm -> "As ==> G" : thm +*) +fun hide_prems th = + let + val sgn = Thm.sign_of_thm th; + val ctermify = Thm.cterm_of sgn; + + val t = (prop_of th); + val prems = Logic.strip_imp_prems t; + val cprems = map ctermify prems; + val aprems = map Thm.trivial cprems; + + (* val unhidef = Drule.implies_intr_list cprems; *) + in + (Drule.implies_elim_list th aprems, cprems) + end; + + + + +(* Fixme: allow different order of subgoals in exportf *) +(* as above, but also fix all parameters in all subgoals, and uses +fix_alls, not fix_alls', ie doesn't leave extra asumptions as apparent +subgoals. *) +(* loosely corresponds to: +Given + "[| !! x0s. A0s x0s ==> SG0 x0s; + ...; !! xms. Ams xms ==> SGm xms|] ==> G" : thm +Result: +(["(A0s x0s' ==> SG0 x0s') ==> SG0 x0s'", + ... ,"(Ams xms' ==> SGm xms') ==> SGm xms'"] : thm list, -- goals + ["SG0 x0s'", ..., "SGm xms'"] : thm list -> -- export function + "G" : thm) +*) +(* requires being given solutions! *) +fun fixed_subgoal_thms th = + let + val (subgoals, expf) = subgoal_thms th; +(* fun export_sg (th, exp) = exp th; *) + fun export_sgs expfs solthms = + expf (map2 (op |>) (solthms, expfs)); +(* expf (map export_sg (ths ~~ expfs)); *) + in + apsnd export_sgs (Library.split_list (map (apsnd export_solution o + fix_alls 1) subgoals)) + end; + +end; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/isaplib.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/isaplib.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,296 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: generic/isaplib.ML + Author: Lucas Dixon, University of Edinburgh + lucasd@dai.ed.ac.uk +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + A few useful system-independent utilities. +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +signature ISAP_LIB = +sig + (* ints *) + val max : int * int -> int + + (* seq operations *) + val ALL_BUT_LAST : ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq + val FST : ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq + val NTH : int -> ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq + val all_but_last_of_seq : 'a Seq.seq -> 'a Seq.seq + val nat_seq : int Seq.seq + val nth_of_seq : int -> 'a Seq.seq -> 'a Library.option + val pair_seq : 'a Seq.seq -> 'b Seq.seq -> ('a * 'b) Seq.seq + val seq_is_empty : 'a Seq.seq -> bool + val number_seq : 'a Seq.seq -> (int * 'a) Seq.seq + + (* lists *) + val mk_num_list : int -> int list + val number_list : int -> 'a list -> (int * 'a) list + val repeated_list : int -> 'a -> 'a list + val all_pairs : 'a list -> 'b list -> ('a * 'b) list + val get_ends_of : ('a * 'a -> bool) -> + ('a * 'a) -> 'a list -> ('a * 'a) + val flatmap : ('a -> 'b list) -> 'a list -> 'b list + + val lrem : ('a * 'b -> bool) -> 'a list -> 'b list -> 'b list + val forall_list : ('a -> bool) -> 'a list -> bool + + (* the list of lists with one of each of the original sublists *) + val one_of_each : 'a list list -> 'a list list + + (* type changing *) + exception NOT_A_DIGIT of string + val int_of_string : string -> int + + (* string manipulations *) + val remove_end_spaces : string -> string + val str_indent : string -> string + val string_of_intlist : int list -> string + val string_of_list : ('a -> string) -> 'a list -> string + + (* options *) + val aptosome : 'a Library.option -> ('a -> 'b) -> 'b Library.option + val seq_mapfilter : ('a -> 'b Library.option) -> 'a Seq.seq -> 'b Seq.seq + val seq_map_to_some_filter : ('a -> 'b) -> 'a Library.option Seq.seq + -> 'b Seq.seq +end; + + + +structure IsaPLib :> ISAP_LIB = +struct + +(* Int *) +fun max (x,y) = if x > y then x else y; + +(* Seq *) +fun seq_map_to_some_filter f s0 = + let + fun recf s () = + case (Seq.pull s) of + None => None + | Some (None,s') => recf s' () + | Some (Some d, s') => + Some (f d, Seq.make (recf s')) + in Seq.make (recf s0) end; + +fun seq_mapfilter f s0 = + let + fun recf s () = + case (Seq.pull s) of + None => None + | Some (a,s') => + (case f a of None => recf s' () + | Some b => Some (b, Seq.make (recf s'))) + in Seq.make (recf s0) end; + + + +(* a simple function to pair with each element of a list, a number *) +fun number_list i [] = [] + | number_list i (h::t) = + (i,h)::(number_list (i+1) t) + +(* check to see if a sequence is empty *) +fun seq_is_empty s = is_none (Seq.pull s); + +(* the sequence of natural numbers *) +val nat_seq = + let fun nseq i () = Some (i, Seq.make (nseq (i+1))) + in Seq.make (nseq 1) + end; + +(* create a sequence of pairs of the elements of the two sequences + If one sequence becomes empty, then so does the pairing of them. + + This is particularly useful if you wish to perform counting or + other repeated operations on a sequence and you want to combvine + this infinite sequence with a possibly finite one. + + behaviour: + s1: a1, a2, ... an + s2: b1, b2, ... bn ... + + pair_seq s1 s2: (a1,b1), (a2,b2), ... (an,bn) +*) +fun pair_seq seq1 seq2 = + let + fun pseq s1 s2 () = + case Seq.pull s1 of + None => None + | Some (s1h, s1t) => + case Seq.pull s2 of + None => None + | Some (s2h, s2t) => + Some ((s1h, s2h), Seq.make (pseq s1t s2t)) + in + Seq.make (pseq seq1 seq2) + end; + +(* number a sequence *) +fun number_seq s = pair_seq nat_seq s; + +(* cuts off the last element of a sequence *) +fun all_but_last_of_seq s = + let + fun f () = + case Seq.pull s of + None => None + | Some (a, s2) => + (case Seq.pull s2 of + None => None + | Some (a2,s3) => + Some (a, all_but_last_of_seq (Seq.cons (a2,s3)))) + in + Seq.make f + end; + + fun ALL_BUT_LAST r st = all_but_last_of_seq (r st); + + + (* nth elem for sequenes, return none if out of bounds *) + fun nth_of_seq i l = + if (seq_is_empty l) then None + else if i <= 1 then Some (Seq.hd l) + else nth_of_seq (i-1) (Seq.tl l); + + (* for use with tactics... gives no_tac if element isn't there *) + fun NTH n f st = + let val r = nth_of_seq n (f st) in + if (is_none r) then Seq.empty else (Seq.single (the r)) + end; + + (* First result of a tactic... uses NTH, so if there is no elements, + + + then no_tac is returned. *) + fun FST f = NTH 1 f; + + (* create a list opf the form (n, n-1, n-2, ... ) *) + fun mk_num_list n = + if n < 1 then [] else (n :: (mk_num_list (n-1))); + + fun repeated_list n a = + if n < 1 then [] else (a :: (repeated_list (n-1) a)); + + (* create all possible pairs with fst element from the first list + and snd element from teh second list *) + fun all_pairs xs ys = + let + fun all_pairs_aux yss [] _ acc = acc + | all_pairs_aux yss (x::xs) [] acc = + all_pairs_aux yss xs yss acc + | all_pairs_aux yss (xs as (x1::x1s)) (y::ys) acc = + all_pairs_aux yss xs ys ((x1,y)::acc) + in + all_pairs_aux ys xs ys [] + end; + + + (* create all possible pairs with fst element from the first list + and snd element from teh second list *) + (* FIXME: make tail recursive and quick *) + fun one_of_each [] = [] + | one_of_each [[]] = [] + | one_of_each [(h::t)] = [h] :: (one_of_each [t]) + | one_of_each ([] :: xs) = [] + | one_of_each ((h :: t) :: xs) = + (map (fn z => h :: z) (one_of_each xs)) + @ (one_of_each (t :: xs)); + + +(* function to get the upper/lower bounds of a list +given: +gr : 'a * 'a -> bool = greater than check +e as (u,l) : ('a * 'a) = upper and lower bits +l : 'a list = the list to get the upper and lower bounds of +returns a pair ('a * 'a) of the biggest and lowest value w.r.t "gr" +*) +fun get_ends_of gr (e as (u,l)) [] = e + | get_ends_of gr (e as (u,l)) (h :: t) = + if gr(h,u) then get_ends_of gr (h,l) t + else if gr(l,h) then get_ends_of gr (u,h) t + else get_ends_of gr (u,l) t; + +fun flatmap f = flat o map f; + + (* quick removal of "rs" elements from "ls" when (eqf (r,l)) is true + Ignores ordering. *) + fun lrem eqf rs ls = + let fun list_remove rs ([],[]) = [] + | list_remove [] (xs,_) = xs + | list_remove (r::rs) ([],leftovers) = + list_remove rs (leftovers,[]) + | list_remove (r1::rs) ((x::xs),leftovers) = + if eqf (r1, x) then list_remove (r1::rs) (xs,leftovers) + else list_remove (r1::rs) (xs, x::leftovers) + in + list_remove rs (ls,[]) + end; + +fun forall_list f [] = true + | forall_list f (a::t) = f a orelse forall_list f t; + + + (* crappy string indeter *) + fun str_indent s = + implode (map (fn s => if s = "\n" then "\n " else s) (explode s)); + + + fun remove_end_spaces s = + let + fun rem_starts [] = [] + | rem_starts (" " :: t) = rem_starts t + | rem_starts ("\t" :: t) = rem_starts t + | rem_starts ("\n" :: t) = rem_starts t + | rem_starts l = l + fun rem_ends l = rev (rem_starts (rev l)) + in + implode (rem_ends (rem_starts (explode s))) + end; + +(* convert a string to an integer *) + exception NOT_A_DIGIT of string; + + fun int_of_string s = + let + fun digits_to_int [] x = x + | digits_to_int (h :: t) x = digits_to_int t (x * 10 + h); + + fun char_to_digit c = + case c of + "0" => 0 + | "1" => 1 + | "2" => 2 + | "3" => 3 + | "4" => 4 + | "5" => 5 + | "6" => 6 + | "7" => 7 + | "8" => 8 + | "9" => 9 + | _ => raise NOT_A_DIGIT c + in + digits_to_int (map char_to_digit (explode (remove_end_spaces s))) 0 + end; + + (* Debugging/printing code for this datatype *) + fun string_of_list f l = + let + fun auxf [] = "" + | auxf [a] = (f a) + | auxf (h :: (t as (h2 :: t2))) = (f h) ^ ", " ^ (auxf t) + in + "[" ^ (auxf l) ^ "]" + end; + + val string_of_intlist = string_of_list string_of_int; + + + (* options *) + fun aptosome None f = None + | aptosome (Some x) f = Some (f x); + +end; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/rw_inst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/rw_inst.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,278 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/rw_inst.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Created: 25 Aug 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + rewriting using a conditional meta-equality theorem which supports + schematic variable instantiation. + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +signature RW_INST = +sig + + (* Rewrite: give it instantiation infromation, a rule, and the + target thm, and it will return the rewritten target thm *) + val rw : + ((Term.indexname * Term.typ) list * (* type var instantiations *) + (Term.indexname * Term.term) list) (* schematic var instantiations *) + * (string * Term.typ) list * (* bound types *) + Term.term -> (* outer term for instantiation *) + Thm.thm -> (* rule with indexies lifted *) + Thm.thm -> (* target thm *) + Thm.thm (* rewritten theorem possibly + with additional premises for + rule conditions *) + + (* used tools *) + val mk_abstractedrule : + (string * Term.typ) list -> Thm.thm -> Thm.cterm list * Thm.thm + val mk_fixtvar_tyinsts : + Term.indexname list -> + Term.term list -> ((string * int) * Term.typ) list * string list + val mk_renamings : + Term.term -> Thm.thm -> (((string * int) * Term.typ) * Term.term) list + val new_tfree : + ((string * int) * Term.sort) * + (((string * int) * Term.typ) list * string list) -> + ((string * int) * Term.typ) list * string list + val cross_inst : (Term.indexname * Term.term) list + -> (Term.indexname * Term.term) list + + val beta_contract_tac : Thm.thm -> Thm.thm + val beta_eta_contract_tac : Thm.thm -> Thm.thm + +end; + +structure RWInst : RW_INST= +struct + + +(* beta contract the theorem *) +fun beta_contract_tac thm = + equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm; + +(* beta-eta contract the theorem *) +fun beta_eta_contract_tac thm = + let + val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm + val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2 + in thm3 end; + +(* Given a list of variables that were bound, and a that has been +instantiated with free variable placeholders for the bound vars, it +creates an abstracted version of the theorem, with local bound vars as +lambda-params: + +Ts: +("x", ty) + +rule:: +C :x ==> P :x = Q :x + +results in: +("!! x. C x", (%x. p x = %y. p y) [!! x. C x]) + +note: assumes rule is instantiated +*) +(* Note, we take abstraction in the order of last abstraction first *) +fun mk_abstractedrule Ts rule = + let + val ctermify = Thm.cterm_of (Thm.sign_of_thm rule); + + (* now we change the names the temporary free vars that represent + bound vars with binders outside the redex *) + val prop = Thm.prop_of rule; + val names = Term.add_term_names (prop, []); + val (fromnames,tonames,names2) = + foldl (fn ((rnf,rnt,names),(n,ty)) => + let val n2 = Term.variant names n in + (ctermify (Free(RWTools.mk_fake_bound_name n, + ty)) :: rnf, + ctermify (Free(n2,ty)) :: rnt, + n2 :: names) + end) + (([],[],names), Ts); + val rule' = rule |> Drule.forall_intr_list fromnames + |> Drule.forall_elim_list tonames; + + (* make unconditional rule and prems *) + val (uncond_rule, cprems) = IsaND.allify_conditions ctermify Ts rule'; + + (* using these names create lambda-abstracted version of the rule *) + val abstractions = Ts ~~ (rev tonames); + val abstract_rule = foldl (fn (th,((n,ty),ct)) => + Thm.abstract_rule n ct th) + (uncond_rule, abstractions); + in (cprems, abstract_rule) end; + + +(* given names to avoid, and vars that need to be fixed, it gives +unique new names to the vars so that they can be fixed as free +variables *) +(* make fixed unique free variable instantiations for non-ground vars *) +(* Create a table of vars to be renamed after instantiation - ie + other uninstantiated vars in the hyps of the rule + ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) +fun mk_renamings tgt rule_inst = + let + val rule_conds = Thm.prems_of rule_inst + val names = foldr Term.add_term_names (tgt :: rule_conds, []); + val (conds_tyvs,cond_vs) = + foldl (fn ((tyvs, vs), t) => + (Library.union + (Term.term_tvars t, tyvs), + Library.union + (map Term.dest_Var (Term.term_vars t), vs))) + (([],[]), rule_conds); + val termvars = map Term.dest_Var (Term.term_vars tgt); + val vars_to_fix = Library.union (termvars, cond_vs); + val (renamings, names2) = + foldr (fn (((n,i),ty), (vs, names')) => + let val n' = Term.variant names' n in + ((((n,i),ty), Free (n', ty)) :: vs, n'::names') + end) + (vars_to_fix, ([], names)); + in renamings end; + +(* make a new fresh typefree instantiation for the given tvar *) +fun new_tfree (tv as (ix,sort), (pairs,used)) = + let val v = variant used (string_of_indexname ix) + in ((ix,TFree(v,sort))::pairs, v::used) end; + + +(* make instantiations to fix type variables that are not + already instantiated (in ignore_ixs) from the list of terms. *) +fun mk_fixtvar_tyinsts ignore_ixs ts = + let val (tvars, tfreenames) = + foldr (fn (t, (varixs, tfreenames)) => + (Term.add_term_tvars (t,varixs), + Term.add_term_tfree_names (t,tfreenames))) + (ts, ([],[])); + val unfixed_tvars = filter (fn (ix,s) => not (ix mem ignore_ixs)) tvars; + val (fixtyinsts, _) = foldr new_tfree (unfixed_tvars, ([], tfreenames)) + in (fixtyinsts, tfreenames) end; + + +(* cross-instantiate the instantiations - ie for each instantiation +replace all occurances in other instantiations - no loops are possible +and thus only one-parsing of the instantiations is necessary. *) +fun cross_inst insts = + let + fun instL (ix, t) = + map (fn (ix2,t2) => (ix2, Term.subst_vars ([], [(ix, t)]) t2)); + + fun cross_instL ([], l) = rev l + | cross_instL ((ix, t) :: insts, l) = + cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); + + in cross_instL (insts, []) end; +(* as above but for types -- I don't know if this is needed, will we ever incur mixed up types? *) +fun cross_inst_typs insts = + let + fun instL (ix, ty) = + map (fn (ix2,ty2) => (ix2, Term.typ_subst_TVars [(ix, ty)] ty2)); + + fun cross_instL ([], l) = rev l + | cross_instL ((ix, t) :: insts, l) = + cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); + + in cross_instL (insts, []) end; + + +(* assume that rule and target_thm have distinct var names *) +(* THINK: efficient version with tables for vars for: target vars, +introduced vars, and rule vars, for quicker instantiation? *) +(* The outerterm defines which part of the target_thm was modified *) +(* Note: we take Ts in the upterm order, ie last abstraction first., +and with an outeterm where the abstracted subterm has the arguments in +the revered order, ie first abstraction first. *) +fun rw ((nonfixed_typinsts, unprepinsts), Ts, outerterm) rule target_thm = + let + (* general signature info *) + val target_sign = (Thm.sign_of_thm target_thm); + val ctermify = Thm.cterm_of target_sign; + val ctypeify = Thm.ctyp_of target_sign; + + (* fix all non-instantiated tvars *) + val (fixtyinsts, othertfrees) = + mk_fixtvar_tyinsts (map fst nonfixed_typinsts) + [Thm.prop_of rule, Thm.prop_of target_thm]; + + val typinsts = cross_inst_typs (nonfixed_typinsts @ fixtyinsts); + + (* certified instantiations for types *) + val ctyp_insts = map (apsnd ctypeify) typinsts; + + (* type instantiated versions *) + val tgt_th_tyinst = + Thm.instantiate (ctyp_insts, []) target_thm; + val rule_tyinst = + Thm.instantiate (ctyp_insts, []) rule; + + (* type instanitated outer term *) + val outerterm_tyinst = + Term.subst_vars (typinsts,[]) outerterm; + + (* type-instantiate the var instantiations *) + val insts_tyinst = foldr (fn ((ix,t),insts_tyinst) => + (ix, Term.subst_vars (typinsts,[]) t) + :: insts_tyinst) + (unprepinsts,[]); + + (* cross-instantiate *) + val insts_tyinst_inst = cross_inst insts_tyinst; + + (* create certms of instantiations *) + val cinsts_tyinst = + map (fn (ix,t) => (ctermify (Var (ix, Term.type_of t)), + ctermify t)) insts_tyinst_inst; + + (* The instantiated rule *) + val rule_inst = rule_tyinst |> Thm.instantiate ([], cinsts_tyinst); + + (* Create a table of vars to be renamed after instantiation - ie + other uninstantiated vars in the hyps the *instantiated* rule + ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) + val renamings = mk_renamings (Thm.prop_of tgt_th_tyinst) + rule_inst; + val cterm_renamings = + map (fn (x,y) => (ctermify (Var x), ctermify y)) renamings; + + (* Create the specific version of the rule for this target application *) + val outerterm_inst = + outerterm_tyinst + |> Term.subst_vars ([], insts_tyinst_inst) + |> Term.subst_vars ([], map (fn ((ix,ty),t) => (ix,t)) + renamings); + val couter_inst = Thm.reflexive (ctermify outerterm_inst); + val (cprems, abstract_rule_inst) = + rule_inst |> Thm.instantiate ([], cterm_renamings) + |> mk_abstractedrule Ts; + val specific_tgt_rule = + beta_eta_contract_tac + (Thm.combination couter_inst abstract_rule_inst); + + (* create an instantiated version of the target thm *) + val tgt_th_inst = + tgt_th_tyinst |> Thm.instantiate ([], cinsts_tyinst) + |> Thm.instantiate ([], cterm_renamings); + + val (vars,frees_of_fixed_vars) = Library.split_list cterm_renamings; + + in + (beta_eta_contract_tac tgt_th_inst) + |> Thm.equal_elim specific_tgt_rule + |> Drule.implies_intr_list cprems + |> Drule.forall_intr_list frees_of_fixed_vars + |> Drule.forall_elim_list vars + |> fst o Thm.varifyT' othertfrees + |> Drule.zero_var_indexes + end; + + +end; (* struct *) diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/rw_tools.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/rw_tools.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,187 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: sys/rw_tools.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Created: 28 May 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + term related tools used for rewriting + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +signature RWTOOLS = +sig +end; + +structure RWTools += struct + +(* THINKABOUT: don't need this: should be able to generate the paired + NsTs directly ? --already implemented as ~~ +fun join_lists ([],[]) = [] + | join_lists (x::xs, y::ys) = (x,y) :: (join_lists (xs,ys)) + | join_lists (_, _) = raise ERROR_MESSAGE "join_lists: unequal size lists"; + *) + +(* fake free variable names for locally bound variables - these work +as placeholders. *) +fun dest_fake_bound_name n = + case (explode n) of + (":" :: realchars) => implode realchars + | _ => n; +fun is_fake_bound_name n = (hd (explode n) = ":"); +fun mk_fake_bound_name n = ":" ^ n; + + + +(* fake free variable names for local meta variables - these work +as placeholders. *) +fun dest_fake_fix_name n = + case (explode n) of + ("@" :: realchars) => implode realchars + | _ => n; +fun is_fake_fix_name n = (hd (explode n) = "@"); +fun mk_fake_fix_name n = "@" ^ n; + + + +(* fake free variable names for meta level bound variables *) +fun dest_fake_all_name n = + case (explode n) of + ("+" :: realchars) => implode realchars + | _ => n; +fun is_fake_all_name n = (hd (explode n) = "+"); +fun mk_fake_all_name n = "+" ^ n; + + + + +(* Ys and Ts not used, Ns are real names of faked local bounds, the +idea is that this will be mapped to free variables thus if a free +variable is a faked local bound then we change it to being a meta +variable so that it can later be instantiated *) +(* FIXME: rename this - avoid the word fix! *) +(* note we are not really "fix"'ing the free, more like making it variable! *) +fun trymkvar_of_fakefree (Ns, Ts) Ys (n,ty) = + if (dest_fake_bound_name n) mem Ns then Var((n,0),ty) else Free (n,ty); + +(* make a var into a fixed free (ie prefixed with "@") *) +fun mk_fakefixvar Ts ((n,i),ty) = Free(mk_fake_fix_name n, ty); + + +(* mk_frees_bound: string list -> Term.term -> Term.term *) +(* This function changes free variables to being represented as bound +variables if the free's variable name is in the given list. The debruijn +index is simply the position in the list *) +(* THINKABOUT: danger of an existing free variable with the same name: fix +this so that name conflict are avoided automatically! In the meantime, +don't have free variables named starting with a ":" *) +fun bounds_of_fakefrees Ys (a $ b) = + (bounds_of_fakefrees Ys a) $ (bounds_of_fakefrees Ys b) + | bounds_of_fakefrees Ys (Abs(n,ty,t)) = + Abs(n,ty, bounds_of_fakefrees (n::Ys) t) + | bounds_of_fakefrees Ys (Free (n,ty)) = + let fun try_mk_bound_of_free (i,[]) = Free (n,ty) + | try_mk_bound_of_free (i,(y::ys)) = + if n = y then Bound i else try_mk_bound_of_free (i+1,ys) + in try_mk_bound_of_free (0,Ys) end + | bounds_of_fakefrees Ys t = t; + + +(* map a function f onto each free variables *) +fun map_to_frees f Ys (a $ b) = + (map_to_frees f Ys a) $ (map_to_frees f Ys b) + | map_to_frees f Ys (Abs(n,ty,t)) = + Abs(n,ty, map_to_frees f ((n,ty)::Ys) t) + | map_to_frees f Ys (Free a) = + f Ys a + | map_to_frees f Ys t = t; + + +(* map a function f onto each meta variable *) +fun map_to_vars f Ys (a $ b) = + (map_to_vars f Ys a) $ (map_to_vars f Ys b) + | map_to_vars f Ys (Abs(n,ty,t)) = + Abs(n,ty, map_to_vars f ((n,ty)::Ys) t) + | map_to_vars f Ys (Var a) = + f Ys a + | map_to_vars f Ys t = t; + +(* map a function f onto each free variables *) +fun map_to_alls f (Const("all",allty) $ Abs(n,ty,t)) = + let val (n2,ty2) = f (n,ty) + in (Const("all",allty) $ Abs(n2,ty2,map_to_alls f t)) end + | map_to_alls f x = x; + +(* map a function f to each type variable in a term *) +(* implicit arg: term *) +fun map_to_term_tvars f = + Term.map_term_types (fn TVar(ix,ty) => f (ix,ty) | x => x); + + + +(* what if a param desn't occur in the concl? think about! Note: This +simply fixes meta level univ bound vars as Frees. At the end, we will +change them back to schematic vars that will then unify +appropriactely, ie with unfake_vars *) +fun fake_concl_of_goal gt i = + let + val prems = Logic.strip_imp_prems gt + val sgt = nth_elem (i - 1, prems) + + val tbody = Logic.strip_imp_concl (Term.strip_all_body sgt) + val tparams = Term.strip_all_vars sgt + + val fakefrees = map (fn (n, ty) => Free(mk_fake_all_name n, ty)) + tparams + in + Term.subst_bounds (rev fakefrees,tbody) + end; + +(* what if a param desn't occur in the concl? think about! Note: This +simply fixes meta level univ bound vars as Frees. At the end, we will +change them back to schematic vars that will then unify +appropriactely, ie with unfake_vars *) +fun fake_goal gt i = + let + val prems = Logic.strip_imp_prems gt + val sgt = nth_elem (i - 1, prems) + + val tbody = Term.strip_all_body sgt + val tparams = Term.strip_all_vars sgt + + val fakefrees = map (fn (n, ty) => Free(mk_fake_all_name n, ty)) + tparams + in + Term.subst_bounds (rev fakefrees,tbody) + end; + + +(* hand written - for some reason the Isabelle version in drule is broken! +Example? something to do with Bin Yangs examples? + *) +fun rename_term_bvars ns (Abs(s,ty,t)) = + let val s2opt = Library.find_first (fn (x,y) => s = x) ns + in case s2opt of + None => (Abs(s,ty,rename_term_bvars ns t)) + | Some (_,s2) => Abs(s2,ty, rename_term_bvars ns t) end + | rename_term_bvars ns (a$b) = + (rename_term_bvars ns a) $ (rename_term_bvars ns b) + | rename_term_bvars _ x = x; + +fun rename_thm_bvars ns th = + let val t = Thm.prop_of th + in Thm.rename_boundvars t (rename_term_bvars ns t) th end; + +(* Finish this to show how it breaks! (raises the exception): + +exception rename_thm_bvars_exp of ((string * string) list * Thm.thm) + + Drule.rename_bvars ns th + handle TERM _ => raise rename_thm_bvars_exp (ns, th); +*) + +end; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/term_lib.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/term_lib.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,824 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: libs/term_lib.ML + Author: Lucas Dixon, University of Edinburgh + lucasd@dai.ed.ac.uk + Created: 17 Aug 2002 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + Additional code to work with terms. + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +signature TERM_LIB = +sig + val fo_term_height : Term.term -> int + val ho_term_height : Term.term -> int + + val current_sign : unit -> Sign.sg + + structure fvartabS : TABLE + + (* Matching/unification with exceptions handled *) + val clean_match : Type.tsig -> Term.term * Term.term + -> ((Term.indexname * Term.typ) list + * (Term.indexname * Term.term) list) Library.option + val clean_unify : Sign.sg -> int -> Term.term * Term.term + -> ((Term.indexname * Term.typ) list + * (Term.indexname * Term.term) list) Seq.seq + + (* managing variables in terms, can doing conversions *) + val bounds_to_frees : Term.term -> Term.term + val bounds_to_frees_with_vars : + (string * Term.typ) list -> Term.term -> Term.term + + val change_bounds_to_frees : Term.term -> Term.term + val change_frees_to_vars : Term.term -> Term.term + val change_vars_to_frees : Term.term -> Term.term + val loose_bnds_to_frees : + (string * Term.typ) list -> Term.term -> Term.term + + (* make all variables named uniquely *) + val unique_namify : Term.term -> Term.term + + (* breasking up a term and putting it into a normal form + independent of internal term context *) + val cleaned_term_conc : Term.term -> Term.term + val cleaned_term_parts : Term.term -> Term.term list * Term.term + val cterm_of_term : Term.term -> Thm.cterm + + (* terms of theorems and subgoals *) + val term_of_thm : Thm.thm -> Term.term + val get_prems_of_sg_term : Term.term -> Term.term list + val triv_thm_from_string : string -> Thm.thm +(* val make_term_into_simp_thm : + (string * Term.typ) list -> Sign.sg -> Term.term -> Thm.thm + val thms_of_prems_in_goal : int -> Thm.thm -> Thm.thm list +*) + + + (* some common term manipulations *) + val try_dest_Goal : Term.term -> Term.term + val try_dest_Trueprop : Term.term -> Term.term + + val bot_left_leaf_of : Term.term -> Term.term + + (* term containing another term - an embedding check *) + val term_contains : Term.term -> Term.term -> bool + + (* name-convertable checking - like ae-convertable, but allows for + var's and free's to be mixed - and doesn't used buggy code. :-) *) + val get_name_eq_member : Term.term -> Term.term list -> Term.term option + val name_eq_member : Term.term -> Term.term list -> bool + val term_name_eq : + Term.term -> + Term.term -> + (Term.term * Term.term) list -> + (Term.term * Term.term) list Library.option + + (* is the typ a function or is it atomic *) + val is_fun_typ : Term.typ -> bool + val is_atomic_typ : Term.typ -> bool + + (* variable analysis *) + val is_some_kind_of_var : Term.term -> bool + val same_var_check : + ''a -> ''b -> (''a * ''b) list -> (''a * ''b) list Library.option + val has_new_vars : Term.term * Term.term -> bool + val has_new_typ_vars : Term.term * Term.term -> bool + (* checks to see if the lhs -> rhs is a invalid rewrite rule *) + val is_not_valid_rwrule : Type.tsig -> (Term.term * Term.term) -> bool + + (* get the frees in a term that are of atomic type, ie non-functions *) + val atomic_frees_of_term : Term.term -> (string * Term.typ) list + + (* Isar term skolemisationm and unsolemisation *) + (* I think this is written also in IsarRTechn and also in somewhere else *) + (* val skolemise_term : (string,Term.typ) list -> Term.term -> Term.term + val unskolemise_term : (string,Term.typ) list -> Term.term -> Term.term *) + val unskolemise_all_term : + Term.term -> + (((string * Term.typ) * string) list * Term.term) + + (* given a string describing term then a string for its type, returns + read term *) + val readwty : string -> string -> Term.term + + (* pretty stuff *) + val pp_term : Term.term -> unit + val pretty_print_sort : string list -> string + val pretty_print_term : Term.term -> string + val pretty_print_type : Term.typ -> string + val pretty_print_typelist : + Term.typ list -> (Term.typ -> string) -> string + + val string_of_term : Term.term -> string + + (* these are perhaps redundent, check the standard basis + lib for generic versions for any datatype? *) + val writesort : string list -> unit + val writeterm : Term.term -> unit + val writetype : Term.typ -> unit + end + + +structure TermLib : TERM_LIB = +struct + +(* Two kinds of depth measure for HOAS terms, a first order (flat) and a + higher order one. + Note: not stable of eta-contraction: embedding eta-expands term, + thus we assume eta-expanded *) +fun fo_term_height (a $ b) = + IsaPLib.max (1 + fo_term_height b, (fo_term_height a)) + | fo_term_height (Abs(_,_,t)) = fo_term_height t + | fo_term_height _ = 0; + +fun ho_term_height (a $ b) = + 1 + (IsaPLib.max (ho_term_height b, ho_term_height a)) + | ho_term_height (Abs(_,_,t)) = ho_term_height t + | ho_term_height _ = 0; + +(* Higher order matching with exception handled *) +(* returns optional instantiation *) +fun clean_match tsig (a as (pat, t)) = + Some (Pattern.match tsig a) handle Pattern.MATCH => None; +(* Higher order unification with exception handled, return the instantiations *) +(* given Signature, max var index, pat, tgt *) +(* returns Seq of instantiations *) +fun clean_unify sgn ix (a as (pat, tgt)) = + let + (* type info will be re-derived, maybe this can be cached + for efficiency? *) + val pat_ty = Term.type_of pat; + val tgt_ty = Term.type_of tgt; + (* is it OK to ignore the type instantiation info? + or should I be using it? *) + val typs_unify = + (Some (Type.unify (Sign.tsig_of sgn) (Term.Vartab.empty, ix) + (pat_ty, tgt_ty))) + handle Type.TUNIFY => None; + in + case typs_unify of + Some (typinsttab, ix2) => + let + (* is it right to throw away teh flexes? + or should I be using them somehow? *) + fun mk_insts (env,flexflexes) = + (Vartab.dest (Envir.type_env env), Envir.alist_of env); + val initenv = Envir.Envir {asol = Vartab.empty, + iTs = typinsttab, maxidx = ix2}; + val useq = (Unify.unifiers (sgn,initenv,[a])) + handle Library.LIST _ => Seq.empty + | Term.TERM _ => Seq.empty; + fun clean_unify' useq () = + (case (Seq.pull useq) of + None => None + | Some (h,t) => Some (mk_insts h, Seq.make (clean_unify' t))) + handle Library.LIST _ => None + | Term.TERM _ => None; + in + (Seq.make (clean_unify' useq)) + end + | None => Seq.empty + end; + +fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); +fun asm_read s = + (Thm.assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); + + +(* more pretty printing code for Isabelle terms etc *) + + +(* pretty_print_typelist l f = print a typelist. + l = list of types to print : typ list + f = function used to print a single type : typ -> string +*) +fun pretty_print_typelist [] f = "" + | pretty_print_typelist [(h: typ)] (f : typ -> string) = (f h) + | pretty_print_typelist ((h: typ) :: t) (f : typ -> string) = + (f h) ^ ", " ^ (pretty_print_typelist t f); + + + +(* pretty_print_sort s = print a sort + s = sort to print : string list +*) +fun pretty_print_sort [] = "" + | pretty_print_sort ([h]) = "\"" ^ h ^ "\"" + | pretty_print_sort (h :: t) = "\"" ^ h ^ "\"," ^ (pretty_print_sort t); + + +(* pretty_print_type t = print a type + t = type to print : type +*) +fun pretty_print_type (Type (n, l)) = + "Type(\"" ^ n ^ "\", [" ^ (pretty_print_typelist l pretty_print_type) ^ "])" + | pretty_print_type (TFree (n, s)) = + "TFree(\"" ^ n ^ "\", [" ^ (pretty_print_sort s) ^ "])" + | pretty_print_type (TVar ((n, i), s)) = + "TVar( (\"" ^ n ^ "\", " ^ (string_of_int i) ^ "), [" ^ (pretty_print_sort s) ^ "])"; + + +(* pretty_print_term t = print a term prints types and sorts too. + t = term to print : term +*) +fun pretty_print_term (Const (s, t)) = + "Const(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" + | pretty_print_term (Free (s, t)) = + "Free(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" + | pretty_print_term (Var ((n, i), t)) = + "Var( (\"" ^ n ^ "\"," ^ (string_of_int i) ^ "), " ^ (pretty_print_type t) ^ ")" + | pretty_print_term (Bound i) = + "Bound(" ^ (string_of_int i) ^ ")" + | pretty_print_term (Abs (s, t, r)) = + "Abs(\"" ^ s ^ "\"," ^ (pretty_print_type t) ^ ", \n " ^ (pretty_print_term r) ^ ")" + | pretty_print_term (op $ (t1, t2)) = + "(" ^ (pretty_print_term t1) ^ ") $\n (" ^ (pretty_print_term t2) ^ ")"; + +(* Write the term out nicly instead of just creating a string for it *) +fun writeterm t = writeln (pretty_print_term t); +fun writetype t = writeln (pretty_print_type t); +fun writesort s = writeln (pretty_print_sort s); + + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Turn on show types *) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +(* if (!show_types) then true else set show_types; *) + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* functions *) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + +fun current_sign () = sign_of (the_context()); +fun cterm_of_term (t : term) = cterm_of (current_sign()) t; +fun term_of_thm t = (Thm.prop_of t); + +(* little function to make a trueprop of anything by putting a P + function in front of it... +fun HOL_mk_term_prop t = + ((Const("Trueprop", Type("fun", + [Type("bool", []), Type("prop", [])]))) $ + ((Free("P", Type("fun", [type_of t, + Type("bool", [])]))) $ t)); + +*) + +fun string_of_term t = + (Sign.string_of_term (current_sign()) t); + +(* +(allt as (Const("Trueprop", ty) $ m)) = + (string_of_cterm (cterm_of_term allt)) + | string_of_term (t : term) = + string_of_cterm (cterm_of_term (HOL_mk_term_prop t)); +*) + +(* creates a simple cterm of the term if it's not already a prop by + prefixing it with a polymorphic function P, then create the cterm + and print that. Handy tool for printing terms that are not + already propositions, or not cterms. +*) +fun pp_term t = writeln (string_of_term t); + +(* create a trivial HOL thm from anything... *) +fun triv_thm_from_string s = + Thm.trivial (cterm_of (current_sign()) (read s)); + + (* Checks if vars could be the same - alpha convertable + w.r.t. previous vars, a and b are assumed to be vars, + free vars, but not bound vars, + Note frees and vars must all have unique names. *) + fun same_var_check a b L = + let + fun bterm_from t [] = None + | bterm_from t ((a,b)::m) = + if t = a then Some b else bterm_from t m + + val bl_opt = bterm_from a L + in + case bterm_from a L of + None => Some ((a,b)::L) + | Some b2 => if b2 = b then Some L else None + end; + + (* FIXME: make more efficient, only require a single new var! *) + (* check if the new term has any meta variables not in the old term *) + fun has_new_vars (old, new) = + (case IsaPLib.lrem (op =) (Term.term_vars old) (Term.term_vars new) of + [] => false + | (_::_) => true); + (* check if the new term has any meta variables not in the old term *) + fun has_new_typ_vars (old, new) = + (case IsaPLib.lrem (op =) (Term.term_tvars old) (Term.term_tvars new) of + [] => false + | (_::_) => true); + + (* working with Isar terms and their skolemisation(s) *) +(* val skolemise_term : (string,Term.typ) list -> Term.term -> Term.term + val unskolemise_term : (string,Term.typ) list -> Term.term -> Term.term +*) + +(* cunning version *) + +(* This version avoids name conflicts that might be intorduced by +unskolemisation, and returns a list of (string * Term.typ) * string, +where the outer string is the original name, and the inner string is +the new name, and the type is the type of the free variable that was +renamed. *) +local + fun myadd (n,ty) (L as (renames, names)) = + let val n' = Syntax.dest_skolem n in + case (Library.find_first (fn (_,n2) => (n2 = n)) renames) of + None => + let val renamedn = Term.variant names n' in + (renamedn, (((renamedn, ty), n) :: renames, + renamedn :: names)) end + | (Some ((renamedn, _), _)) => (renamedn, L) + end + handle LIST _ => (n, L); + + fun unsk (L,f) (t1 $ t2) = + let val (L', t1') = unsk (L, I) t1 + in unsk (L', (fn x => f (t1' $ x))) t2 end + | unsk (L,f) (Abs(n,ty,t)) = + unsk (L, (fn x => Abs(n,ty,x))) t + | unsk (L, f) (Free (n,ty)) = + let val (renamed_n, L') = myadd (n ,ty) L + in (L', f (Free (renamed_n, ty))) end + | unsk (L, f) l = (L, f l); +in +fun unskolemise_all_term t = + let + val names = Term.add_term_names (t,[]) + val ((renames,names'),t') = unsk (([], names),I) t + in (renames,t') end; +end + +(* fun unskolemise_all_term t = + let + fun fmap fv = + let (n,ty) = (Term.dest_Free fv) in + Some (fv, Free (Syntax.dest_skolem n, ty)) handle LIST _ => None + end + val substfrees = mapfilter fmap (Term.term_frees t) + in + Term.subst_free substfrees t + end; *) + +(* true if the type t is a function *) +fun is_fun_typ (Type(s, l)) = + if s = "fun" then true else false + | is_fun_typ _ = false; + +val is_atomic_typ = not o is_fun_typ; + +(* get the frees in a term that are of atomic type, ie non-functions *) +val atomic_frees_of_term = + filter (is_atomic_typ o snd) + o map Term.dest_Free + o Term.term_frees; + + +(* read in a string and a top-level type and this will give back a term *) +fun readwty tstr tystr = + let + val sgn = Theory.sign_of (the_context()) + in + Sign.simple_read_term sgn (Sign.read_typ (sgn, K None) tystr) tstr + end; + + + (* first term is equal to the second in some name convertable + way... Note: This differs from the aeconv in the Term.ML file of + Isabelle in that it allows a var to be alpha-equiv to a free var. + + Also, the isabelle term.ML version of aeconv uses a + function that it claims doesn't work! *) + + fun term_name_eq (Abs(_,ty1,t1)) (Abs(_,ty2,t2)) l = + if ty1 = ty2 then term_name_eq t1 t2 l + else None + + | term_name_eq (ah $ at) (bh $ bt) l = + let + val lopt = (term_name_eq ah bh l) + in + if is_some lopt then + term_name_eq at bt (the lopt) + else + None + end + + | term_name_eq (Const(a,T)) (Const(b,U)) l = + if a=b andalso T=U then Some l + else None + + | term_name_eq (a as Free(s1,ty1)) (b as Free(s2,ty2)) l = + same_var_check a b l + + | term_name_eq (a as Free(s1,ty1)) (b as Var(n2,ty2)) l = + same_var_check a b l + + | term_name_eq (a as Var(n1,ty1)) (b as Free(s2,ty2)) l = + same_var_check a b l + + | term_name_eq (a as Var(n1,ty1)) (b as Var(n2,ty2)) l = + same_var_check a b l + + | term_name_eq (Bound i) (Bound j) l = + if i = j then Some l else None + + | term_name_eq a b l = ((*writeln ("unchecked case:\n" ^ "a:\n" ^ (pretty_print_term a) ^ "\nb:\n" ^ (pretty_print_term b));*) None); + + (* checks to see if the term in a name-equivalent member of the list of terms. *) + fun get_name_eq_member a [] = None + | get_name_eq_member a (h :: t) = + if is_some (term_name_eq a h []) then Some h + else get_name_eq_member a t; + + fun name_eq_member a [] = false + | name_eq_member a (h :: t) = + if is_some (term_name_eq a h []) then true + else name_eq_member a t; + + (* true if term is a variable *) + fun is_some_kind_of_var (Free(s, ty)) = true + | is_some_kind_of_var (Var(i, ty)) = true + | is_some_kind_of_var (Bound(i)) = true + | is_some_kind_of_var _ = false; + + (* checks to see if the lhs -> rhs is a invalid rewrite rule *) +(* FIXME: we should really check that there is a subterm on the lhs +which embeds into the rhs, this would be much closer to the normal +notion of valid wave rule - ie there exists at least one case where it +is a valid wave rule... *) + fun is_not_valid_rwrule tysig (lhs, rhs) = + (Pattern.matches_subterm tysig (lhs, rhs)) orelse + has_new_vars (lhs,rhs) orelse + has_new_typ_vars (lhs,rhs); + + + (* first term contains the second in some name convertable way... *) + (* note: this is equivalent to an alpha-convertable emedding *) + (* takes as args: term containing a, term contained b, + (bound vars of a, bound vars of b), + current alpha-conversion-pairs, + returns: bool:can convert, new alpha-conversion table *) + (* in bellow: we *don't* use: a loose notion that only requires + variables to match variables, and doesn't worry about the actual + pattern in the variables. *) + fun term_contains1 (Bs, FVs) (Abs(s,ty,t)) (Abs(s2,ty2,t2)) = + if ty = ty2 then + term_contains1 ((Some(s,s2,ty)::Bs), FVs) t t2 + else [] + + | term_contains1 T t1 (Abs(s2,ty2,t2)) = [] + + | term_contains1 (Bs, FVs) (Abs(s,ty,t)) t2 = + term_contains1 (None::Bs, FVs) t t2 + + | term_contains1 T (ah $ at) (bh $ bt) = + (term_contains1 T ah (bh $ bt)) @ + (term_contains1 T at (bh $ bt)) @ + (flat (map (fn inT => (term_contains1 inT at bt)) + (term_contains1 T ah bh))) + + | term_contains1 T a (bh $ bt) = [] + + | term_contains1 T (ah $ at) b = + (term_contains1 T ah b) @ (term_contains1 T at b) + + | term_contains1 T a b = + (* simple list table lookup to check if a named variable has been + mapped to a variable, if not adds the mapping and return some new + list mapping, if it is then it checks that the pair are mapped to + each other, if so returns the current mapping list, else none. *) + let + fun bterm_from t [] = None + | bterm_from t ((a,b)::m) = + if t = a then Some b else bterm_from t m + + (* check to see if, w.r.t. the variable mapping, two terms are leaf + terms and are mapped to each other. Note constants are only mapped + to the same constant. *) + fun same_leaf_check (T as (Bs,FVs)) (Bound i) (Bound j) = + let + fun aux_chk (i1,i2) [] = false + | aux_chk (0,0) ((Some _) :: bnds) = true + | aux_chk (i1,0) (None :: bnds) = false + | aux_chk (i1,i2) ((Some _) :: bnds) = + aux_chk (i1 - 1,i2 - 1) bnds + | aux_chk (i1,i2) (None :: bnds) = + aux_chk (i1,i2 - 1) bnds + in + if (aux_chk (i,j) Bs) then [T] + else [] + end + | same_leaf_check (T as (Bs,(Fs,Vs))) + (a as (Free (an,aty))) + (b as (Free (bn,bty))) = + (case bterm_from an Fs of + Some b2n => if bn = b2n then [T] + else [] (* conflict of var name *) + | None => [(Bs,((an,bn)::Fs,Vs))]) + | same_leaf_check (T as (Bs,(Fs,Vs))) + (a as (Var (an,aty))) + (b as (Var (bn,bty))) = + (case bterm_from an Vs of + Some b2n => if bn = b2n then [T] + else [] (* conflict of var name *) + | None => [(Bs,(Fs,(an,bn)::Vs))]) + | same_leaf_check T (a as (Const _)) (b as (Const _)) = + if a = b then [T] else [] + | same_leaf_check T _ _ = [] + + in + same_leaf_check T a b + end; + + (* wrapper for term_contains1: checks if the term "a" contains in + some embedded way, (w.r.t. name -convertable) the term "b" *) + fun term_contains a b = + case term_contains1 ([],([],[])) a b of + (_ :: _) => true + | [] => false; + + + (* pp_term a; pp_term b; writeln "EQTerm matches are: "; map (fn (a,b) => writeln ("(" ^ (string_of_term a) ^ ", " ^ (string_of_term b) ^ ")")) L;*) + + + (* change all bound variables to see ones with appropriate name and + type *) + + (* naming convention is OK as we use 'variant' from term.ML *) + + (* Note "bounds_to_frees" defined below, its better and quicker, but + keeps the quantifiers handing about, and changes all bounds, not + just universally quantified ones. *) + fun change_bounds_to_frees t = + let + val vars = strip_all_vars t + val frees_names = map (fn Free(s,n) => s | _ => "") (term_frees t) + val body = strip_all_body t + + fun bnds_to_frees [] _ acc = acc + | bnds_to_frees ((name,ty)::more) names acc = + let + val new_name = variant names name + in + bnds_to_frees more (new_name::names) (Free(new_name,ty)::acc) + end + in + (subst_bounds ((bnds_to_frees vars frees_names []), body)) + end; + + + +structure fvartabS = TableFun(type key = string val ord = string_ord); + +(* a runtime-quick function which makes sure that every variable has a +unique name *) +fun unique_namify_aux (ntab,(Abs(s,ty,t))) = + (case (fvartabS.lookup (ntab,s)) of + None => + let + val (ntab2,t2) = unique_namify_aux ((fvartabS.update ((s,s),ntab)), t) + in + (ntab2,Abs(s,ty,t2)) + end + | Some s2 => + let + val s_new = (Term.variant (fvartabS.keys ntab) s) + val (ntab2,t2) = + unique_namify_aux ((fvartabS.update ((s_new,s_new),ntab)), t) + in + (ntab2,Abs(s_new,ty,t2)) + end) + | unique_namify_aux (ntab,(a $ b)) = + let + val (ntab1,t1) = unique_namify_aux (ntab,a) + val (ntab2,t2) = unique_namify_aux (ntab1,b) + in + (ntab2, t1$t2) + end + | unique_namify_aux (nt as (ntab,Const x)) = nt + | unique_namify_aux (nt as (ntab,f as Free (s,ty))) = + (case (fvartabS.lookup (ntab,s)) of + None => ((fvartabS.update ((s,s),ntab)), f) + | Some _ => nt) + | unique_namify_aux (nt as (ntab,v as Var ((s,i),ty))) = + (case (fvartabS.lookup (ntab,s)) of + None => ((fvartabS.update ((s,s),ntab)), v) + | Some _ => nt) + | unique_namify_aux (nt as (ntab, Bound i)) = nt; + +fun unique_namify t = + #2 (unique_namify_aux (fvartabS.empty, t)); + +(* a runtime-quick function which makes sure that every variable has a +unique name and also changes bound variables to free variables, used +for embedding checks, Note that this is a pretty naughty term +manipulation, which doesn't have necessary relation to the original +sematics of the term. This is really a trick for our embedding code. *) + +fun bounds_to_frees_aux T (ntab,(Abs(s,ty,t))) = + (case (fvartabS.lookup (ntab,s)) of + None => + let + val (ntab2,t2) = bounds_to_frees_aux ((s,ty)::T) + ((fvartabS.update ((s,s),ntab)), t) + in + (ntab2,Abs(s,ty,t2)) + end + | Some s2 => + let + val s_new = (Term.variant (fvartabS.keys ntab) s) + val (ntab2,t2) = + bounds_to_frees_aux ((s_new,ty)::T) + (fvartabS.update (((s_new,s_new),ntab)), t) + in + (ntab2,Abs(s_new,ty,t2)) + end) + | bounds_to_frees_aux T (ntab,(a $ b)) = + let + val (ntab1,t1) = bounds_to_frees_aux T (ntab,a) + val (ntab2,t2) = bounds_to_frees_aux T (ntab1,b) + in + (ntab2, t1$t2) + end + | bounds_to_frees_aux T (nt as (ntab,Const x)) = nt + | bounds_to_frees_aux T (nt as (ntab,f as Free (s,ty))) = + (case (fvartabS.lookup (ntab,s)) of + None => ((fvartabS.update ((s,s),ntab)), f) + | Some _ => nt) + | bounds_to_frees_aux T (nt as (ntab,v as Var ((s,i),ty))) = + (case (fvartabS.lookup (ntab,s)) of + None => ((fvartabS.update ((s,s),ntab)), v) + | Some _ => nt) + | bounds_to_frees_aux T (nt as (ntab, Bound i)) = + let + val (s,ty) = Library.nth_elem (i,T) + in + (ntab, Free (s,ty)) + end; + + +fun bounds_to_frees t = + #2 (bounds_to_frees_aux [] (fvartabS.empty,t)); + +fun bounds_to_frees_with_vars vars t = + #2 (bounds_to_frees_aux vars (fvartabS.empty,t)); + + + +(* loose bounds to frees *) +fun loose_bnds_to_frees_aux (bnds,vars) (Abs(s,ty,t)) = + Abs(s,ty,loose_bnds_to_frees_aux (bnds + 1,vars) t) + | loose_bnds_to_frees_aux (bnds,vars) (a $ b) = + (loose_bnds_to_frees_aux (bnds,vars) a) + $ (loose_bnds_to_frees_aux (bnds,vars) b) + | loose_bnds_to_frees_aux (bnds,vars) (t as (Bound i)) = + if (bnds <= i) then Free (Library.nth_elem (i - bnds,vars)) + else t + | loose_bnds_to_frees_aux (bnds,vars) t = t; + + +fun loose_bnds_to_frees vars t = + loose_bnds_to_frees_aux (0,vars) t; + +(* puts prems of a theorem into a useful form, (rulify) + Note that any loose bound variables are treated as free vars +*) +(* fun make_term_into_simp_thm vars sgn t = + let + val triv = + Thm.trivial (Thm.cterm_of + sgn (loose_bnds_to_frees vars t)) + in + SplitterData.mk_eq (Drule.standard (ObjectLogic.rulify_no_asm triv)) + end; + +fun thms_of_prems_in_goal i tm= + let + val goal = (nth_elem (i-1,Thm.prems_of tm)) + val vars = rev (strip_all_vars goal) + val prems = Logic.strip_assums_hyp (strip_all_body goal) + val simp_prem_thms = + map (make_term_into_simp_thm vars (Thm.sign_of_thm tm)) prems + in + simp_prem_thms + end; +*) + + (* replace all universally quantifief variables (at HOL object level) + with free vars + fun HOL_Alls_to_frees TL (Const("All", _) $ Abs(v, ty, t)) = + HOL_Alls_to_frees ((Free (v, ty)) :: TL) t + | HOL_Alls_to_frees TL t = + (TL,(subst_bounds (TL, t))); +*) + + (* this is just a hack for mixing with interactive mode, and using topthm() *) + + fun try_dest_Goal (Const("Goal", _) $ T) = T + | try_dest_Goal T = T; + + fun try_dest_Trueprop (Const("Trueprop", _) $ T) = T + | try_dest_Trueprop T = T; + + fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l + | bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t + | bot_left_leaf_of x = x; + + +(* cleaned up normal form version of the subgoal terms conclusion *) +fun cleaned_term_conc t = + let + val concl = Logic.strip_imp_concl (change_bounds_to_frees t) + in + (try_dest_Trueprop (try_dest_Goal concl)) + end; + +(* fun get_prems_of_sg_term t = + map opt_dest_Trueprop (Logic.strip_imp_prems t); *) + +fun get_prems_of_sg_term t = + map try_dest_Trueprop (Logic.strip_assums_hyp t); + + +(* drop premices, clean bound var stuff, and make a trueprop... *) + fun cleaned_term_parts t = + let + val t2 = (change_bounds_to_frees t) + val concl = Logic.strip_imp_concl t2 + val prems = map try_dest_Trueprop (Logic.strip_imp_prems t2) + in + (prems, (try_dest_Trueprop (try_dest_Goal concl))) + end; + +(* old version + fun cleaned_term t = + let + val concl = (HOLogic.dest_Trueprop (Logic.strip_imp_concl + (change_bounds_to_frees t) )) + in + concl + end; +*) + + (* change free variables to vars and visa versa *) + (* *** check naming is OK, can we just use the vasr old name? *) + (* *** Check: Logic.varify and Logic.unvarify *) + fun change_vars_to_frees (a$b) = + (change_vars_to_frees a) $ (change_vars_to_frees b) + | change_vars_to_frees (Abs(s,ty,t)) = + (Abs(s,Type.varifyT ty,change_vars_to_frees t)) + | change_vars_to_frees (Var((s,i),ty)) = (Free(s,Type.unvarifyT ty)) + | change_vars_to_frees l = l; + + fun change_frees_to_vars (a$b) = + (change_frees_to_vars a) $ (change_frees_to_vars b) + | change_frees_to_vars (Abs(s,ty,t)) = + (Abs(s,Type.varifyT ty,change_frees_to_vars t)) + | change_frees_to_vars (Free(s,ty)) = (Var((s,0),Type.varifyT ty)) + | change_frees_to_vars l = l; + + +end; + + + + (* ignores lambda abstractions, ie (\x y = y) + same as embedding check code, ie. (f a b) = can "a" be embedded in "b" + *) + (* fun term_is_same_or_simpler_than (Abs(s,ty,t)) b = + term_is_same_or_simpler_than t b + + | term_is_same_or_simpler_than a (Abs(s,ty,t)) = + term_is_same_or_simpler_than a t + + | term_is_same_or_simpler_than (ah $ at) (bh $ bt) = + (term_is_same_or_simpler_than (ah $ at) bh) orelse + (term_is_same_or_simpler_than (ah $ at) bt) orelse + ((term_is_same_or_simpler_than ah bh) andalso + (term_is_same_or_simpler_than at bt)) + + | term_is_same_or_simpler_than (ah $ at) b = false + + | term_is_same_or_simpler_than a (bh $ bt) = + (term_is_same_or_simpler_than a bh) orelse + (term_is_same_or_simpler_than a bt) + + | term_is_same_or_simpler_than a b = + if a = b then true else false; + *) + +(* fun term_is_simpler_than t1 t2 = + (are_different_terms t1 t2) andalso + (term_is_same_or_simpler_than t1 t2); +*) diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/IsaPlanner/upterm_lib.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/IsaPlanner/upterm_lib.ML Tue Feb 01 18:01:57 2005 +0100 @@ -0,0 +1,132 @@ +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* Title: libs/upterm_lib.ML + Author: Lucas Dixon, University of Edinburgh + lucas.dixon@ed.ac.uk + Created: 26 Jan 2004 +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +(* DESCRIPTION: + + generic upterms for lambda calculus + +*) +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) + + +signature UPTERM_LIB = +sig + + + (* type for upterms defined in terms of a 't : a downterm type, and + 'y : types for bound variable in the downterm type *) + datatype ('t,'y) T = + abs of string * 'y * (('t,'y) T) + | appl of 't * (('t,'y) T) + | appr of 't * (('t,'y) T) + | root + + (* analysis *) + val tyenv_of_upterm : ('t,'y) T -> (string * 'y) list + val tyenv_of_upterm' : ('t,'y) T -> 'y list + val addr_of_upterm : ('t,'y) T -> int list + val upsize_of_upterm : ('t,'y) T -> int + + (* editing *) + val map_to_upterm_parts : ('t -> 't2) * ('y -> 'y2) -> + ('t,'y) T -> ('t2,'y2) T + + val expand_map_to_upterm_parts : ('t -> 't2 list) * ('y -> 'y2) -> + ('t,'y) T -> ('t2,'y2) T list + + val fold_upterm : ('a * 't -> 'a) -> (* left *) + ('a * 't -> 'a) -> (* right *) + ('a * (string * 'y) -> 'a) -> (* abs *) + ('a * ('t,'y) T) -> 'a (* everything *) + + (* apply one term to another *) + val apply : ('t,'y) T -> ('t,'y) T -> ('t,'y) T + +end; + + + +(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) +structure UpTermLib : UPTERM_LIB = +struct + + (* type for upterms defined in terms of a 't : a downterm type, and + 'y : types for bound variable in the downterm type *) + datatype ('t,'y) T = + abs of string * 'y * ('t,'y) T + | appl of 't * ('t,'y) T + | appr of 't * ('t,'y) T + | root; + + (* get the bound variable names and types for the current foucs *) + fun tyenv_of_upterm (appl(l,m)) = tyenv_of_upterm m + | tyenv_of_upterm (appr(r,m)) = tyenv_of_upterm m + | tyenv_of_upterm (abs(s,ty,m)) = (s, ty) :: (tyenv_of_upterm m) + | tyenv_of_upterm root = []; + + (* get the bound variable types for the current foucs *) + fun tyenv_of_upterm' (appl(l,m)) = tyenv_of_upterm' m + | tyenv_of_upterm' (appr(r,m)) = tyenv_of_upterm' m + | tyenv_of_upterm' (abs(s,ty,m)) = ty :: (tyenv_of_upterm' m) + | tyenv_of_upterm' root = []; + + (* an address construction for the upterm, ie if we were at the + root, address describes how to get to this point. *) + fun addr_of_upterm1 A root = A + | addr_of_upterm1 A (appl (l,m)) = addr_of_upterm1 (1::A) m + | addr_of_upterm1 A (appr (r,m)) = addr_of_upterm1 (2::A) m + | addr_of_upterm1 A (abs (s,ty,m)) = addr_of_upterm1 (0::A) m; + + fun addr_of_upterm m = addr_of_upterm1 [] m; + + (* the size of the upterm, ie how far to get to root *) + fun upsize_of_upterm root = 0 + | upsize_of_upterm (appl (l,m)) = (upsize_of_upterm m) + 1 + | upsize_of_upterm (appr (r,m)) = (upsize_of_upterm m) + 1 + | upsize_of_upterm (abs (s,ty,m)) = (upsize_of_upterm m) + 1; + + (* apply an upterm to to another upterm *) + fun apply x root = x + | apply x (appl (l,m)) = appl(l, apply x m) + | apply x (appr (r,m)) = appr(r, apply x m) + | apply x (abs (s,ty,m)) = abs(s, ty, apply x m); + + (* applies the term function to each term in each part of the upterm *) + fun map_to_upterm_parts (tf,yf) root = root + + | map_to_upterm_parts (tf,yf) (abs(s,ty,m)) = + abs(s,yf ty,map_to_upterm_parts (tf,yf) m) + + | map_to_upterm_parts (tf,yf) (appr(t,m)) = + appr (tf t, map_to_upterm_parts (tf,yf) m) + + | map_to_upterm_parts (tf,yf) (appl(t,m)) = + appl (tf t, map_to_upterm_parts (tf,yf) m); + + + (* applies the term function to each term in each part of the upterm *) + fun expand_map_to_upterm_parts (tf,yf) root = [root] + | expand_map_to_upterm_parts (tf,yf) (abs(s,ty,m)) = + map (fn x => abs(s,yf ty, x)) + (expand_map_to_upterm_parts (tf,yf) m) + | expand_map_to_upterm_parts (tf,yf) (appr(t,m)) = + map appr (IsaPLib.all_pairs + (tf t) (expand_map_to_upterm_parts (tf,yf) m)) + | expand_map_to_upterm_parts (tf,yf) (appl(t,m)) = + map appl (IsaPLib.all_pairs + (tf t) (expand_map_to_upterm_parts (tf,yf) m)); + + (* fold along each 't (down term) in the upterm *) + fun fold_upterm fl fr fa (d, u) = + let + fun fold_upterm' (d, root) = d + | fold_upterm' (d, abs(s,ty,m)) = fold_upterm' (fa(d,(s,ty)), m) + | fold_upterm' (d, appr(t,m)) = fold_upterm' (fr(d,t), m) + | fold_upterm' (d, appl(t,m)) = fold_upterm' (fl(d,t), m) + in fold_upterm' (d,u) end; + +end; diff -r cb3612cc41a3 -r fc075ae929e4 src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Sun Jan 30 20:48:50 2005 +0100 +++ b/src/Pure/ROOT.ML Tue Feb 01 18:01:57 2005 +0100 @@ -66,6 +66,9 @@ (*old-style goal package*) use "goals.ML"; +(*the IsaPlanner subsystem*) +cd "IsaPlanner"; use "ROOT.ML"; cd ".."; + (*configuration for Proof General*) use "proof_general.ML"; diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/AC/HH.thy Tue Feb 01 18:01:57 2005 +0100 @@ -16,10 +16,7 @@ "HH(f,x,a) == transrec(a, %b r. let z = x - (\c \ b. r`c) in if f`z \ Pow(z)-{0} then f`z else {x})" - -(* ********************************************************************** *) -(* Lemmas useful in each of the three proofs *) -(* ********************************************************************** *) +subsection{*Lemmas useful in each of the three proofs*} lemma HH_def_satisfies_eq: "HH(f,x,a) = (let z = x - (\b \ a. HH(f,x,b)) @@ -42,14 +39,12 @@ apply (fast intro!: ltI intro: Ord_in_Ord) done -lemma Diff_UN_eq_self: - "(!!y. y \ A ==> P(y) = {x}) ==> x - (\y \ A. P(y)) = x" -apply (simp, fast elim!: mem_irrefl) -done +lemma Diff_UN_eq_self: "(!!y. y\A ==> P(y) = {x}) ==> x - (\y \ A. P(y)) = x" +by (simp, fast elim!: mem_irrefl) lemma HH_eq: "x - (\b \ a. HH(f,x,b)) = x - (\b \ a1. HH(f,x,b)) ==> HH(f,x,a) = HH(f,x,a1)" -apply (subst HH_def_satisfies_eq) +apply (subst HH_def_satisfies_eq [of _ _ a1]) apply (rule HH_def_satisfies_eq [THEN trans], simp) done @@ -131,9 +126,7 @@ apply (erule_tac [2] ltI [OF _ Ord_Least], assumption) done -(* ********************************************************************** *) -(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) -(* ********************************************************************** *) +subsection{*Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1*} lemma lam_Least_HH_inj_Pow: "(\a \ (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) @@ -220,9 +213,7 @@ lam_sing_bij [THEN bij_converse_bij], standard] -(* ********************************************************************** *) -(* The proof of AC1 ==> WO2 *) -(* ********************************************************************** *) +subsection{*The proof of AC1 ==> WO2*} (*Establishing the existence of a bijection, namely converse diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/AC/WO2_AC16.thy Tue Feb 01 18:01:57 2005 +0100 @@ -440,8 +440,8 @@ (\T \ recfunAC16(f, h, y, a). ~ h`b \ T))" apply (drule ex_next_set, assumption+) apply (erule bexE) -apply (rule oexI) -apply (subst right_inverse_bij, blast, assumption+) +apply (rule_tac x="converse(f)`X" in oexI) +apply (simp add: right_inverse_bij) apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI Card_is_Ord) done diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/ArithSimp.thy --- a/src/ZF/ArithSimp.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/ArithSimp.thy Tue Feb 01 18:01:57 2005 +0100 @@ -7,8 +7,13 @@ header{*Arithmetic with simplification*} -theory ArithSimp = Arith -files "arith_data.ML": +theory ArithSimp +imports Arith +files "~~/src/Provers/Arith/cancel_numerals.ML" + "~~/src/Provers/Arith/combine_numerals.ML" + "arith_data.ML" + +begin subsection{*Difference*} diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Constructible/Formula.thy --- a/src/ZF/Constructible/Formula.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Constructible/Formula.thy Tue Feb 01 18:01:57 2005 +0100 @@ -633,21 +633,26 @@ text{*Kunen's VI 1.6 (b)*} lemma Lset_mono [rule_format]: "ALL j. i<=j --> Lset(i) <= Lset(j)" -apply (rule_tac a=i in eps_induct) -apply (rule impI [THEN allI]) -apply (subst Lset) -apply (subst Lset, blast) -done +proof (induct i rule: eps_induct, intro allI impI) + fix x j + assume "\y\x. \j. y \ j \ Lset(y) \ Lset(j)" + and "x \ j" + thus "Lset(x) \ Lset(j)" + by (force simp add: Lset [of x] Lset [of j]) +qed text{*This version lets us remove the premise @{term "Ord(i)"} sometimes.*} lemma Lset_mono_mem [rule_format]: "ALL j. i:j --> Lset(i) <= Lset(j)" -apply (rule_tac a=i in eps_induct) -apply (rule impI [THEN allI]) -apply (subst Lset, auto) -apply (rule rev_bexI, assumption) -apply (blast intro: elem_subset_in_DPow dest: LsetD DPowD) -done +proof (induct i rule: eps_induct, intro allI impI) + fix x j + assume "\y\x. \j. y \ j \ Lset(y) \ Lset(j)" + and "x \ j" + thus "Lset(x) \ Lset(j)" + by (force simp add: Lset [of j] + intro!: bexI intro: elem_subset_in_DPow dest: LsetD DPowD) +qed + text{*Useful with Reflection to bump up the ordinal*} lemma subset_Lset_ltD: "[|A \ Lset(i); i < j|] ==> A \ Lset(j)" diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Epsilon.thy Tue Feb 01 18:01:57 2005 +0100 @@ -246,8 +246,7 @@ lemma rank_mono: "a<=b ==> rank(a) le rank(b)" apply (rule subset_imp_le) -apply (subst rank) -apply (subst rank, auto) +apply (auto simp add: rank [of a] rank [of b]) done lemma rank_Pow: "rank(Pow(a)) = succ(rank(a))" diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Induct/Multiset.thy Tue Feb 01 18:01:57 2005 +0100 @@ -787,7 +787,7 @@ M0 \ acc(multirel1(A, r)); a \ A; \M. \ multirel1(A, r) --> M +# {#a#} \ acc(multirel1(A, r)) |] ==> M0 +# {#a#} \ acc(multirel1(A, r))" -apply (subgoal_tac "M0 \ Mult (A) ") +apply (subgoal_tac "M0 \ Mult(A) ") prefer 2 apply (erule acc.cases) apply (erule fieldE) @@ -1000,7 +1000,7 @@ apply (erule_tac P = "\k \ mset_of (K) . ?P (k) " in rev_mp) apply (erule ssubst) apply (simp add: Ball_def, auto) -apply (subgoal_tac "< (I +# {# x \ K. \ r#}) +# {# x \ K. \ r#}, (I +# {# x \ K. \ r#}) +# J'> \ multirel (A, r) ") +apply (subgoal_tac "< (I +# {# x \ K. \ r#}) +# {# x \ K. \ r#}, (I +# {# x \ K. \ r#}) +# J'> \ multirel(A, r) ") prefer 2 apply (drule_tac x = "I +# {# x \ K. \ r#}" in spec) apply (rotate_tac -1) @@ -1100,7 +1100,7 @@ apply (erule rev_mp) apply (erule trancl_induct, clarify) apply (blast intro: munion_multirel1_mono r_into_trancl, clarify) -apply (subgoal_tac "y \ Mult (A) ") +apply (subgoal_tac "y \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [unfolded multirel_def, THEN subsetD]) apply (subgoal_tac " \ multirel1 (A, r) ") @@ -1111,8 +1111,8 @@ lemma munion_multirel_mono1: "[| \ multirel(A, r); K \ Mult(A)|] ==> \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) -apply (rule_tac P = "%x. \ multirel (A, r) " in munion_commute [THEN subst]) -apply (subst munion_commute [symmetric]) +apply (rule_tac P = "%x. \ multirel(A, r) " in munion_commute [THEN subst]) +apply (subst munion_commute [of N]) apply (rule munion_multirel_mono2) apply (auto simp add: Mult_iff_multiset) done @@ -1120,7 +1120,7 @@ lemma munion_multirel_mono: "[| \ multirel(A, r); \ multirel(A, r)|] ==> \ multirel(A, r)" -apply (subgoal_tac "M \ Mult (A) & N \ Mult (A) & K \ Mult (A) & L \ Mult (A) ") +apply (subgoal_tac "M \ Mult(A) & N \ Mult(A) & K \ Mult(A) & L \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [THEN subsetD]) apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2) done @@ -1272,10 +1272,10 @@ lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M" apply (simp add: mle_def mless_def) -apply (subgoal_tac "\i. Ord (i) & M \ Mult (field (Memrel (i))) ") +apply (subgoal_tac "\i. Ord (i) & M \ Mult(field(Memrel(i))) ") prefer 2 apply (simp add: omultiset_def) apply (case_tac "M=0", simp_all, clarify) -apply (subgoal_tac "<0 +# 0, 0 +# M> \ multirel (field (Memrel (i)), Memrel (i))") +apply (subgoal_tac "<0 +# 0, 0 +# M> \ multirel(field (Memrel(i)), Memrel(i))") apply (rule_tac [2] one_step_implies_multirel) apply (auto simp add: Mult_iff_multiset) done diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Integ/IntDiv.thy --- a/src/ZF/Integ/IntDiv.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Integ/IntDiv.thy Tue Feb 01 18:01:57 2005 +0100 @@ -1204,7 +1204,7 @@ apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") prefer 2 apply (simp add: zcompare_rls) apply (simp (no_asm_simp) add: zadd_zmult_distrib2) -apply (subst zadd_commute, rule zadd_zless_mono) +apply (subst zadd_commute [of "b $\ q'"], rule zadd_zless_mono) prefer 2 apply (blast intro: zmult_zle_mono1) apply (subgoal_tac "r' $+ #0 $< b $+ r") apply (simp add: zcompare_rls) diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/ROOT.ML --- a/src/ZF/ROOT.ML Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/ROOT.ML Tue Feb 01 18:01:57 2005 +0100 @@ -18,9 +18,6 @@ (*syntax for old-style theory sections*) use "thy_syntax"; -use "~~/src/Provers/Arith/cancel_numerals.ML"; -use "~~/src/Provers/Arith/combine_numerals.ML"; - with_path "Integ" use_thy "Main_ZFC"; print_depth 8; diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Resid/Residuals.thy --- a/src/ZF/Resid/Residuals.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Resid/Residuals.thy Tue Feb 01 18:01:57 2005 +0100 @@ -34,9 +34,7 @@ res_func_def: "u |> v == THE w. residuals(u,v,w)" -(* ------------------------------------------------------------------------- *) -(* Setting up rule lists *) -(* ------------------------------------------------------------------------- *) +subsection{*Setting up rule lists*} declare Sres.intros [intro] declare Sreg.intros [intro] @@ -67,10 +65,7 @@ declare Sres.intros [simp] - -(* ------------------------------------------------------------------------- *) -(* residuals is a partial function *) -(* ------------------------------------------------------------------------- *) +subsection{*residuals is a partial function*} lemma residuals_function [rule_format]: "residuals(u,v,w) ==> \w1. residuals(u,v,w1) --> w1 = w" @@ -87,10 +82,7 @@ apply (blast intro: residuals_function)+ done - -(* ------------------------------------------------------------------------- *) -(* Residual function *) -(* ------------------------------------------------------------------------- *) +subsection{*Residual function*} lemma res_Var [simp]: "n \ nat ==> Var(n) |> Var(n) = Var(n)" by (unfold res_func_def, blast) @@ -120,9 +112,7 @@ "[|s~t; regular(t)|]==> regular(t) --> s |> t \ redexes" by (erule Scomp.induct, auto) -(* ------------------------------------------------------------------------- *) -(* Commutation theorem *) -(* ------------------------------------------------------------------------- *) +subsection{*Commutation theorem*} lemma sub_comp [simp]: "u<==v ==> u~v" by (erule Ssub.induct, simp_all) @@ -154,9 +144,7 @@ by (simp add: residuals_subst_rec) -(* ------------------------------------------------------------------------- *) -(* Residuals are comp and regular *) -(* ------------------------------------------------------------------------- *) +subsection{*Residuals are comp and regular*} lemma residuals_preserve_comp [rule_format, simp]: "u~v ==> \w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)" @@ -167,9 +155,7 @@ apply (erule Scomp.induct, auto) done -(* ------------------------------------------------------------------------- *) -(* Preservation lemma *) -(* ------------------------------------------------------------------------- *) +subsection{*Preservation lemma*} lemma union_preserve_comp: "u~v ==> v ~ (u un v)" by (erule Scomp.induct, simp_all) @@ -182,39 +168,30 @@ done -(**** And now the Cube ***) - declare sub_comp [THEN comp_sym, simp] -(* ------------------------------------------------------------------------- *) -(* Prism theorem *) -(* ============= *) -(* ------------------------------------------------------------------------- *) +subsection{*Prism theorem*} (* Having more assumptions than needed -- removed below *) lemma prism_l [rule_format]: "v<==u ==> regular(u) --> (\w. w~v --> w~u --> w |> u = (w|>v) |> (u|>v))" -apply (erule Ssub.induct, force+) -done +by (erule Ssub.induct, force+) -lemma prism: - "[|v <== u; regular(u); w~v|] ==> w |> u = (w|>v) |> (u|>v)" +lemma prism: "[|v <== u; regular(u); w~v|] ==> w |> u = (w|>v) |> (u|>v)" apply (rule prism_l) apply (rule_tac [4] comp_trans, auto) done -(* ------------------------------------------------------------------------- *) -(* Levy's Cube Lemma *) -(* ------------------------------------------------------------------------- *) +subsection{*Levy's Cube Lemma*} lemma cube: "[|u~v; regular(v); regular(u); w~u|]==> (w|>u) |> (v|>u) = (w|>v) |> (u|>v)" -apply (subst preservation, assumption, assumption) -apply (subst preservation, erule comp_sym, assumption) -apply (subst prism [symmetric]) +apply (subst preservation [of u], assumption, assumption) +apply (subst preservation [of v], erule comp_sym, assumption) +apply (subst prism [symmetric, of v]) apply (simp add: union_r comp_sym_iff) apply (simp add: union_preserve_regular comp_sym_iff) apply (erule comp_trans, assumption) @@ -223,9 +200,7 @@ done -(* ------------------------------------------------------------------------- *) -(* paving theorem *) -(* ------------------------------------------------------------------------- *) +subsection{*paving theorem*} lemma paving: "[|w~u; w~v; regular(u); regular(v)|]==> \uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u)~vu & diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/UNITY/AllocImpl.thy --- a/src/ZF/UNITY/AllocImpl.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/UNITY/AllocImpl.thy Tue Feb 01 18:01:57 2005 +0100 @@ -284,8 +284,10 @@ apply (rule_tac [2] mono_length) prefer 3 apply simp apply (simp_all add: refl_prefix Le_def comp_def length_type) -apply (subst Int_commute) -apply (rule_tac A = " ({s \ state . k \ length (s ` rel) } \ {s\state . s ` NbR = n}) \ {s\state. k \ length (s`rel) }" in LeadsTo_weaken_L) +apply (subst Int_commute [of _ "{x \ state . n < x ` NbR}"]) +apply (rule_tac A = "({s \ state . k \ length (s ` rel) } \ + {s\state . s ` NbR = n}) \ {s\state. k \ length(s`rel)}" + in LeadsTo_weaken_L) apply (rule PSP_Stable, safe) apply (rule_tac B = "{x \ state . n < length (x ` rel) } \ {s\state . s ` NbR = n}" in LeadsTo_Trans) apply (rule_tac [2] LeadsTo_weaken) diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/UNITY/FP.thy --- a/src/ZF/UNITY/FP.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/UNITY/FP.thy Tue Feb 01 18:01:57 2005 +0100 @@ -37,15 +37,13 @@ done lemma stable_FP_Orig_Int: "F \ program ==> F \ stable(FP_Orig(F) Int B)" -apply (unfold FP_Orig_def stable_def) -apply (subst Int_Union2) +apply (simp only: FP_Orig_def stable_def Int_Union2) apply (blast intro: constrains_UN) done lemma FP_Orig_weakest2: "[| \B. F \ stable (A Int B); st_set(A) |] ==> A \ FP_Orig(F)" -apply (unfold FP_Orig_def stable_def st_set_def, blast) -done +by (unfold FP_Orig_def stable_def st_set_def, blast) lemmas FP_Orig_weakest = allI [THEN FP_Orig_weakest2, standard] diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/Univ.thy --- a/src/ZF/Univ.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/Univ.thy Tue Feb 01 18:01:57 2005 +0100 @@ -47,8 +47,8 @@ "A<=B ==> \j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)" apply (rule_tac a=i in eps_induct) apply (rule impI [THEN allI]) -apply (subst Vfrom) -apply (subst Vfrom) +apply (subst Vfrom [of A]) +apply (subst Vfrom [of B]) apply (erule Un_mono) apply (erule UN_mono, blast) done @@ -59,18 +59,21 @@ subsubsection{* A fundamental equality: Vfrom does not require ordinals! *} + + lemma Vfrom_rank_subset1: "Vfrom(A,x) <= Vfrom(A,rank(x))" -apply (rule_tac a=x in eps_induct) -apply (subst Vfrom) -apply (subst Vfrom) -apply (blast intro!: rank_lt [THEN ltD]) -done +proof (induct x rule: eps_induct) + fix x + assume "\y\x. Vfrom(A,y) \ Vfrom(A,rank(y))" + thus "Vfrom(A, x) \ Vfrom(A, rank(x))" + by (simp add: Vfrom [of _ x] Vfrom [of _ "rank(x)"], + blast intro!: rank_lt [THEN ltD]) +qed lemma Vfrom_rank_subset2: "Vfrom(A,rank(x)) <= Vfrom(A,x)" apply (rule_tac a=x in eps_induct) apply (subst Vfrom) -apply (subst Vfrom) -apply (rule subset_refl [THEN Un_mono]) +apply (subst Vfrom, rule subset_refl [THEN Un_mono]) apply (rule UN_least) txt{*expand @{text "rank(x1) = (\y\x1. succ(rank(y)))"} in assumptions*} apply (erule rank [THEN equalityD1, THEN subsetD, THEN UN_E]) diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/ZF.thy --- a/src/ZF/ZF.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/ZF.thy Tue Feb 01 18:01:57 2005 +0100 @@ -310,6 +310,8 @@ lemma ballI [intro!]: "[| !!x. x\A ==> P(x) |] ==> \x\A. P(x)" by (simp add: Ball_def) +lemmas strip = impI allI ballI + lemma bspec [dest?]: "[| \x\A. P(x); x: A |] ==> P(x)" by (simp add: Ball_def) diff -r cb3612cc41a3 -r fc075ae929e4 src/ZF/ex/Limit.thy --- a/src/ZF/ex/Limit.thy Sun Jan 30 20:48:50 2005 +0100 +++ b/src/ZF/ex/Limit.thy Tue Feb 01 18:01:57 2005 +0100 @@ -1242,13 +1242,13 @@ apply (erule chain_mkcpo) done -lemma subcpo_Dinf: - "emb_chain(DD,ee) ==> subcpo(Dinf(DD,ee),iprod(DD))" +lemma subcpo_Dinf: "emb_chain(DD,ee) ==> subcpo(Dinf(DD,ee),iprod(DD))" apply (simp add: Dinf_def) apply (rule subcpo_mkcpo) apply (simp add: Dinf_def [symmetric]) apply (rule ballI) -apply (subst lub_iprod) +apply (simplesubst lub_iprod) + --{*Subst would rewrite the lhs. We want to change the rhs.*} apply (assumption | rule chain_Dinf emb_chain_cpo)+ apply simp apply (subst Rp_cont [THEN cont_lub]) @@ -1688,7 +1688,8 @@ apply (rule_tac i = "succ (na) " and j = n in nat_linear_le) apply blast apply assumption - apply (subst eps_split_right_le) + apply (simplesubst eps_split_right_le) + --{*Subst would rewrite the lhs. We want to change the rhs.*} prefer 2 apply assumption apply simp apply (assumption | rule add_le_self nat_0I nat_succI)+