src/HOL/Tools/Function/function_core.ML
changeset 80926 eabc517e7413
parent 80925 6c1146e6e79e
equal deleted inserted replaced
80925:6c1146e6e79e 80926:eabc517e7413
   385   in
   385   in
   386     (exactly_one, function_value)
   386     (exactly_one, function_value)
   387   end
   387   end
   388 
   388 
   389 
   389 
   390 fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
   390 fun prove_stuff ctxt globals G0 f R0 clauses complete compat compat_store G_elim f_def =
   391   let
   391   let
   392     val Globals {h, domT, ranT, x, ...} = globals
   392     val Globals {h, domT, ranT, x = x0, ...} = globals
       
   393 
       
   394     val G = Thm.cterm_of ctxt G0
       
   395     val R = Thm.cterm_of ctxt R0
       
   396 
       
   397     val x = Thm.cterm_of ctxt x0
       
   398 
       
   399     val A = Thm.ctyp_of ctxt domT
       
   400     val B = Thm.ctyp_of ctxt ranT
       
   401     val C = Thm.ctyp_of_cterm x
   393 
   402 
   394     val ihyp =
   403     val ihyp =
   395       \<^instantiate>\<open>'a = \<open>Thm.ctyp_of ctxt domT\<close>
   404       \<^instantiate>\<open>'a = A and 'b = B and 'c = C and x and R and G
   396         and 'b = \<open>Thm.ctyp_of ctxt ranT\<close>
       
   397         and 'c = \<open>Thm.ctyp_of ctxt (fastype_of x)\<close>
       
   398         and x = \<open>Thm.cterm_of ctxt x\<close>
       
   399         and R = \<open>Thm.cterm_of ctxt R\<close>
       
   400         and G = \<open>Thm.cterm_of ctxt G\<close>
       
   401         in cprop \<open>\<And>z::'a. R z x \<Longrightarrow> \<exists>!y::'b. G z y\<close> for x :: 'c\<close>
   405         in cprop \<open>\<And>z::'a. R z x \<Longrightarrow> \<exists>!y::'b. G z y\<close> for x :: 'c\<close>
       
   406 
       
   407     val acc_R_x =
       
   408       \<^instantiate>\<open>'c = C and R and x
       
   409         in cprop \<open>Wellfounded.accp R x\<close> for x :: 'c\<close>
   402 
   410 
   403     val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
   411     val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
   404     val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
   412     val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
   405     val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
   413     val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
   406       |> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)]
   414       |> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)]
   410 
   418 
   411     val _ = trace_msg (K "Proving cases for unique existence...")
   419     val _ = trace_msg (K "Proving cases for unique existence...")
   412     val (ex1s, values) =
   420     val (ex1s, values) =
   413       split_list
   421       split_list
   414         (map
   422         (map
   415           (mk_uniqueness_case ctxt globals G f ihyp ih_intro G_elim compat_store clauses repLemmas)
   423           (mk_uniqueness_case ctxt globals G0 f ihyp ih_intro G_elim compat_store clauses repLemmas)
   416           clauses)
   424           clauses)
   417 
   425 
   418     val _ = trace_msg (K "Proving: Graph is a function")
   426     val _ = trace_msg (K "Proving: Graph is a function")
   419     val graph_is_function = complete
   427     val graph_is_function = complete
   420       |> Thm.forall_elim_vars 0
   428       |> Thm.forall_elim_vars 0
   421       |> fold (curry op COMP) ex1s
   429       |> fold (curry op COMP) ex1s
   422       |> Thm.implies_intr (ihyp)
   430       |> Thm.implies_intr ihyp
   423       |> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
   431       |> Thm.implies_intr acc_R_x
   424       |> Thm.forall_intr (Thm.cterm_of ctxt x)
   432       |> Thm.forall_intr x
   425       |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
   433       |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
   426       |> (fn it =>
   434       |> (fn it =>
   427           fold (Thm.forall_intr o Thm.cterm_of ctxt o Var)
   435           fold (Thm.forall_intr o Thm.cterm_of ctxt o Var)
   428             (Term.add_vars (Thm.prop_of it) []) it)
   436             (Term.add_vars (Thm.prop_of it) []) it)
   429 
   437 
   768   in
   776   in
   769     Function_Context_Tree.traverse_tree step tree
   777     Function_Context_Tree.traverse_tree step tree
   770   end
   778   end
   771 
   779 
   772 
   780 
   773 fun mk_nest_term_rule ctxt globals R R_cases clauses =
   781 fun mk_nest_term_rule ctxt globals R0 R_cases clauses =
   774   let
   782   let
   775     val Globals { domT, x, z, ... } = globals
   783     val Globals { domT, x = x0, z = z0, ... } = globals
   776     val acc_R = mk_acc domT R
       
   777 
   784 
   778     val ([Rn], ctxt') = Variable.variant_fixes ["R"] ctxt
   785     val ([Rn], ctxt') = Variable.variant_fixes ["R"] ctxt
   779     val R' = Free (Rn, fastype_of R)
   786 
   780 
   787     val R = Thm.cterm_of ctxt' R0
   781     val Rrel = Free (Rn, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
   788     val R' = Thm.cterm_of ctxt' (Free (Rn, Thm.typ_of_cterm R))
   782     val inrel_R = \<^Const>\<open>in_rel domT domT for Rrel\<close>
   789     val Rrel = Thm.cterm_of ctxt' (Free (Rn, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))))
   783 
   790 
   784     val wfR' =
   791     val x = Thm.cterm_of ctxt' x0
   785       \<^instantiate>\<open>'a = \<open>Thm.ctyp_of ctxt' domT\<close>
   792     val z = Thm.cterm_of ctxt' z0
   786         and R = \<open>Thm.cterm_of ctxt' R'\<close> in cprop \<open>wfP R\<close>\<close>
   793 
   787 
   794     val A = Thm.ctyp_of ctxt' domT
   788     (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
   795     val B = Thm.ctyp_of_cterm x
   789     val ihyp = Logic.all_const domT $ Abs ("z", domT,
   796 
   790       Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
   797     val acc_R_z =
   791         HOLogic.mk_Trueprop (acc_R $ Bound 0)))
   798       \<^instantiate>\<open>'a = A and R and z
   792       |> Thm.cterm_of ctxt'
   799         in cterm \<open>Wellfounded.accp R z\<close> for z :: 'a\<close>
       
   800 
       
   801     val inrel_R =
       
   802       \<^instantiate>\<open>'a = A and Rrel
       
   803         in cterm \<open>in_rel Rrel\<close> for Rrel :: \<open>('a \<times> 'a) set\<close>\<close>
       
   804 
       
   805     val wfR' = \<^instantiate>\<open>'a = A and R' in cprop \<open>wfP R'\<close>\<close>
       
   806 
       
   807     val ihyp =
       
   808       \<^instantiate>\<open>'a = A and 'b = B and x and R' and R
       
   809         in cprop \<open>\<And>z::'a. R' z x \<Longrightarrow> Wellfounded.accp R z\<close> for x :: 'b\<close>
   793 
   810 
   794     val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
   811     val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
   795 
   812 
   796     val R_z_x = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (R $ z $ x))
   813     val R_z_x =
   797 
   814       \<^instantiate>\<open>'a = A and 'b = B and R and x and z
   798     val (hyps, cases) = fold (mk_nest_term_case ctxt' globals R' ihyp_a) clauses ([], [])
   815         in cprop \<open>R z x\<close> for x :: 'a and z :: 'b\<close>
       
   816 
       
   817     val (hyps, cases) =
       
   818       fold (mk_nest_term_case ctxt' globals (Thm.term_of R') ihyp_a) clauses ([], [])
   799   in
   819   in
   800     R_cases
   820     R_cases
   801     |> Thm.forall_elim (Thm.cterm_of ctxt' z)
   821     |> Thm.forall_elim z
   802     |> Thm.forall_elim (Thm.cterm_of ctxt' x)
   822     |> Thm.forall_elim x
   803     |> Thm.forall_elim (Thm.cterm_of ctxt' (acc_R $ z))
   823     |> Thm.forall_elim acc_R_z
   804     |> curry op COMP (Thm.assume R_z_x)
   824     |> curry op COMP (Thm.assume R_z_x)
   805     |> fold_rev (curry op COMP) cases
   825     |> fold_rev (curry op COMP) cases
   806     |> Thm.implies_intr R_z_x
   826     |> Thm.implies_intr R_z_x
   807     |> Thm.forall_intr (Thm.cterm_of ctxt' z)
   827     |> Thm.forall_intr z
   808     |> (fn it => it COMP accI)
   828     |> (fn it => it COMP accI)
   809     |> Thm.implies_intr ihyp
   829     |> Thm.implies_intr ihyp
   810     |> Thm.forall_intr (Thm.cterm_of ctxt' x)
   830     |> Thm.forall_intr x
   811     |> (fn it => Drule.compose (it, 2, wf_induct_rule))
   831     |> (fn it => Drule.compose (it, 2, wf_induct_rule))
   812     |> curry op RS (Thm.assume wfR')
   832     |> curry op RS (Thm.assume wfR')
   813     |> Thm.forall_intr_vars
   833     |> Thm.forall_intr_vars
   814     |> (fn it => it COMP allI)
   834     |> (fn it => it COMP allI)
   815     |> fold Thm.implies_intr hyps
   835     |> fold Thm.implies_intr hyps
   816     |> Thm.implies_intr wfR'
   836     |> Thm.implies_intr wfR'
   817     |> Thm.forall_intr (Thm.cterm_of ctxt' R')
   837     |> Thm.forall_intr R'
   818     |> Thm.forall_elim (Thm.cterm_of ctxt' inrel_R)
   838     |> Thm.forall_elim inrel_R
   819     |> curry op RS wf_in_rel
   839     |> curry op RS wf_in_rel
   820     |> full_simplify (put_simpset HOL_basic_ss ctxt' addsimps [in_rel_def])
   840     |> full_simplify (put_simpset HOL_basic_ss ctxt' addsimps [in_rel_def])
   821     |> Thm.forall_intr_name ("R", Thm.cterm_of ctxt' Rrel)
   841     |> Thm.forall_intr_name ("R", Rrel)
   822   end
   842   end
   823 
   843 
   824 
   844 
   825 
   845 
   826 fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
   846 fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =