src/HOL/Tools/function_package/fundef_proof.ML
author krauss
Wed, 20 Sep 2006 12:05:31 +0200
changeset 20638 241792a4634e
parent 20523 36a59e5d0039
child 21051 c49467a9c1e1
permissions -rw-r--r--
Removed "induct set" attribute from total induction rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/fundef_proof.ML
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    ID:         $Id$
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     4
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     5
A package for general recursive function definitions. 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     6
Internal proofs.
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     8
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     9
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    10
signature FUNDEF_PROOF =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    11
sig
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    12
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    13
    val mk_partial_rules : theory -> FundefCommon.prep_result 
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    14
			   -> thm -> FundefCommon.fundef_result
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    15
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    16
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    17
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    18
structure FundefProof : FUNDEF_PROOF = 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    19
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    20
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    21
open FundefCommon
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    22
open FundefAbbrev
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    23
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    24
(* Theory dependencies *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    25
val subsetD = thm "subsetD"
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
val subset_refl = thm "subset_refl"
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    27
val split_apply = thm "Product_Type.split"
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    28
val wf_induct_rule = thm "wf_induct_rule";
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
val Pair_inject = thm "Product_Type.Pair_inject";
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    30
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    31
val acc_induct_rule = thm "Accessible_Part.acc_induct_rule"
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    32
val acc_downward = thm "Accessible_Part.acc_downward"
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    33
val accI = thm "Accessible_Part.accI"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    35
val acc_subset_induct = thm "Accessible_Part.acc_subset_induct"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    37
val conjunctionD1 = thm "conjunctionD1"
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    38
val conjunctionD2 = thm "conjunctionD2"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    40
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    41
fun mk_psimp thy globals R f_iff graph_is_function clause valthm =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    42
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    43
	val Globals {domT, z, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    44
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    45
	val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {qs, cqs, gs, lhs, rhs, ags, ...}, ...} = clause
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    46
	val lhs_acc = cterm_of thy (Trueprop (mk_mem (lhs, mk_acc domT R))) (* "lhs : acc R" *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    47
	val z_smaller = cterm_of thy (Trueprop (mk_relmemT domT domT (z, lhs) R)) (* "(z, lhs) : R" *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    48
    in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    49
	((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    50
	    |> (fn it => it COMP graph_is_function)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    51
	    |> implies_intr z_smaller
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    52
	    |> forall_intr (cterm_of thy z)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    53
	    |> (fn it => it COMP valthm)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    54
	    |> implies_intr lhs_acc 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    55
	    |> asm_simplify (HOL_basic_ss addsimps [f_iff])
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    56
            |> fold_rev (implies_intr o cprop_of) ags
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    57
            |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    58
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    59
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    60
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    61
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    62
fun mk_partial_induct_rule thy globals R complete_thm clauses =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    63
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    64
	val Globals {domT, x, z, a, P, D, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    65
        val acc_R = mk_acc domT R
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    66
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    67
	val x_D = assume (cterm_of thy (Trueprop (mk_mem (x, D))))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
	val a_D = cterm_of thy (Trueprop (mk_mem (a, D)))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    69
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    70
	val D_subset = cterm_of thy (Trueprop (mk_subset domT D acc_R))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    71
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    72
	val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    73
	    mk_forall x
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    74
		      (mk_forall z (Logic.mk_implies (Trueprop (mk_mem (x, D)),
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    75
						      Logic.mk_implies (mk_relmem (z, x) R,
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    76
									Trueprop (mk_mem (z, D))))))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    77
		      |> cterm_of thy
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    78
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    79
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    80
	(* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    81
	val ihyp = all domT $ Abs ("z", domT, 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    82
				   implies $ Trueprop (mk_relmemT domT domT (Bound 0, x) R)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    83
					   $ Trueprop (P $ Bound 0))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    84
		       |> cterm_of thy
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    85
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    86
	val aihyp = assume ihyp
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    87
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    88
	fun prove_case clause =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    89
	    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    90
		val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, gs, lhs, rhs, case_hyp, ...}, RCs, 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
    91
                                qglr = (oqs, _, _, _), ...} = clause
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    92
								       
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    93
		val replace_x_ss = HOL_basic_ss addsimps [case_hyp]
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    94
		val lhs_D = simplify replace_x_ss x_D (* lhs : D *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    95
		val sih = full_simplify replace_x_ss aihyp
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    96
					
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    97
                fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    98
                    sih |> forall_elim (cterm_of thy rcarg)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    99
                        |> implies_elim_swp llRI
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   100
                        |> fold_rev (implies_intr o cprop_of) CCas
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   101
                        |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   102
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   103
                val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   104
			     
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   105
		val step = Trueprop (P $ lhs)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   106
				    |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   107
				    |> fold_rev (curry Logic.mk_implies) gs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   108
				    |> curry Logic.mk_implies (Trueprop (mk_mem (lhs, D)))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   109
				    |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   110
				    |> cterm_of thy
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   111
			   
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   112
		val P_lhs = assume step
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   113
				   |> fold forall_elim cqs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   114
				   |> implies_elim_swp lhs_D 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   115
				   |> fold_rev implies_elim_swp ags
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   116
				   |> fold implies_elim_swp P_recs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   117
			    
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   118
		val res = cterm_of thy (Trueprop (P $ x))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   119
				   |> Simplifier.rewrite replace_x_ss
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   120
				   |> symmetric (* P lhs == P x *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   121
				   |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   122
				   |> implies_intr (cprop_of case_hyp)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   123
				   |> fold_rev (implies_intr o cprop_of) ags
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   124
				   |> fold_rev forall_intr cqs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   125
	    in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   126
		(res, step)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   127
	    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   128
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   129
	val (cases, steps) = split_list (map prove_case clauses)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   130
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   131
	val istep =  complete_thm
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   132
                       |> forall_elim_vars 0
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   133
		       |> fold (curry op COMP) cases (*  P x  *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   134
		       |> implies_intr ihyp
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   135
		       |> implies_intr (cprop_of x_D)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   136
		       |> forall_intr (cterm_of thy x)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   137
			   
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   138
	val subset_induct_rule = 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   139
	    acc_subset_induct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   140
		|> (curry op COMP) (assume D_subset)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   141
		|> (curry op COMP) (assume D_dcl)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   142
		|> (curry op COMP) (assume a_D)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   143
		|> (curry op COMP) istep
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   144
		|> fold_rev implies_intr steps
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   145
		|> implies_intr a_D
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   146
		|> implies_intr D_dcl
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   147
		|> implies_intr D_subset
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   148
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   149
	val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   150
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   151
	val simple_induct_rule =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   152
	    subset_induct_rule
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   153
		|> forall_intr (cterm_of thy D)
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   154
		|> forall_elim (cterm_of thy acc_R)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   155
		|> (curry op COMP) subset_refl
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   156
		|> (curry op COMP) (acc_downward
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   157
					|> (instantiate' [SOME (ctyp_of thy domT)]
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   158
							 (map (SOME o cterm_of thy) [x, R, z]))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   159
					|> forall_intr (cterm_of thy z)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   160
					|> forall_intr (cterm_of thy x))
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   161
		|> forall_intr (cterm_of thy a)
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   162
		|> forall_intr (cterm_of thy P)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   163
    in
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   164
	(subset_induct_all, simple_induct_rule)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   165
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   166
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   167
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   168
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   169
(* Does this work with Guards??? *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   170
fun mk_domain_intro thy globals R R_cases clause =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   171
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   172
	val Globals {z, domT, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   173
	val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...}, 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   174
                        qglr = (oqs, _, _, _), ...} = clause
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   175
	val goal = (HOLogic.mk_Trueprop (HOLogic.mk_mem (lhs, mk_acc domT R)))
19782
48c4632e2c28 HOL/Tools/function_package: imporoved handling of guards, added an example
krauss
parents: 19781
diff changeset
   176
                     |> fold_rev (curry Logic.mk_implies) gs
48c4632e2c28 HOL/Tools/function_package: imporoved handling of guards, added an example
krauss
parents: 19781
diff changeset
   177
                     |> cterm_of thy
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   178
    in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   179
	Goal.init goal 
19781
c62720b20e9a HOL/Tools/function_package: More cleanup
krauss
parents: 19773
diff changeset
   180
		  |> (SINGLE (resolve_tac [accI] 1)) |> the
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   181
		  |> (SINGLE (eresolve_tac [forall_elim_vars 0 R_cases] 1))  |> the
19781
c62720b20e9a HOL/Tools/function_package: More cleanup
krauss
parents: 19773
diff changeset
   182
		  |> (SINGLE (CLASIMPSET auto_tac)) |> the
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   183
		  |> Goal.conclude
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   184
                  |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   185
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   186
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   187
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   188
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   189
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   190
fun mk_nest_term_case thy globals R' ihyp clause =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   191
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   192
	val Globals {x, z, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   193
	val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   194
                        qglr=(oqs, _, _, _), ...} = clause
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   195
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   196
	val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   197
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   198
	fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) = 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   199
	    let
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   200
		val used = map (fn ((f,a),thm) => FundefCtxTree.export_thm thy (f, map prop_of a) thm) (u @ sub)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   201
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   202
		val hyp = mk_relmem (arg, lhs) R'
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   203
				    |> fold_rev (curry Logic.mk_implies o prop_of) used
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   204
				    |> FundefCtxTree.export_term (fixes, map prop_of assumes) 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   205
				    |> fold_rev (curry Logic.mk_implies o prop_of) ags
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   206
				    |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   207
				    |> cterm_of thy
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   208
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   209
		val thm = assume hyp
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   210
				 |> fold forall_elim cqs
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   211
				 |> fold implies_elim_swp ags
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   212
				 |> FundefCtxTree.import_thm thy (fixes, assumes) (*  "(arg, lhs) : R'"  *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   213
				 |> fold implies_elim_swp used
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   214
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   215
		val acc = thm COMP ih_case
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   216
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   217
		val z_eq_arg = cterm_of thy (Trueprop (HOLogic.mk_eq (z, arg)))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   218
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   219
		val arg_eq_z = (assume z_eq_arg) RS sym
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   220
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   221
		val z_acc = simplify (HOL_basic_ss addsimps [arg_eq_z]) acc (* fragile, slow... *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   222
				     |> implies_intr (cprop_of case_hyp)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   223
				     |> implies_intr z_eq_arg
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   224
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   225
		val zx_eq_arg_lhs = cterm_of thy (Trueprop (mk_eq (mk_prod (z,x), mk_prod (arg,lhs))))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   226
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   227
		val ethm = (z_acc COMP (assume zx_eq_arg_lhs COMP Pair_inject))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   228
			       |> FundefCtxTree.export_thm thy (fixes, 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   229
                                                                (term_of zx_eq_arg_lhs) :: map prop_of (ags @ assumes))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   230
                               |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   231
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   232
		val sub' = sub @ [(([],[]), acc)]
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   233
	    in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   234
		(sub', (hyp :: hyps, ethm :: thms))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   235
	    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   236
	  | step _ _ _ _ = raise Match
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   237
    in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   238
	FundefCtxTree.traverse_tree step tree
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   239
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   240
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   241
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   242
fun mk_nest_term_rule thy globals R R_cases clauses =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   243
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   244
	val Globals { domT, x, z, ... } = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   245
        val acc_R = mk_acc domT R
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   246
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   247
	val R' = Free ("R", fastype_of R)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   248
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   249
	val wfR' = cterm_of thy (Trueprop (Const ("Wellfounded_Recursion.wf", mk_relT (domT, domT) --> boolT) $ R')) (* "wf R'" *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   250
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   251
	(* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   252
	val ihyp = all domT $ Abs ("z", domT, 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   253
				   implies $ Trueprop (mk_relmemT domT domT (Bound 0, x) R')
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   254
					   $ Trueprop ((Const ("op :", [domT, HOLogic.mk_setT domT] ---> boolT))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   255
							   $ Bound 0 $ acc_R))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   256
		       |> cterm_of thy
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   257
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   258
	val ihyp_a = assume ihyp |> forall_elim_vars 0
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   259
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   260
	val z_ltR_x = cterm_of thy (mk_relmem (z, x) R) (* "(z, x) : R" *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   261
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   262
	val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   263
    in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   264
	R_cases
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   265
            |> forall_elim (cterm_of thy (mk_prod (z,x)))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   266
            |> forall_elim (cterm_of thy (mk_mem (z, acc_R)))
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   267
	    |> curry op COMP (assume z_ltR_x)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   268
	    |> fold_rev (curry op COMP) cases
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   269
	    |> implies_intr z_ltR_x
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   270
	    |> forall_intr (cterm_of thy z)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   271
	    |> (fn it => it COMP accI)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   272
	    |> implies_intr ihyp
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   273
	    |> forall_intr (cterm_of thy x)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   274
	    |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   275
	    |> implies_elim_swp (assume wfR')
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   276
	    |> fold implies_intr hyps
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   277
	    |> implies_intr wfR'
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   278
	    |> forall_intr (cterm_of thy R')
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   279
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   280
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   281
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   282
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   283
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   284
fun mk_partial_rules thy data provedgoal =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   285
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   286
	val Prep {globals, G, f, R, clauses, values, R_cases, ex1_iff, ...} = data
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   287
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   288
        val _ = print "Closing Derivation"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   289
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   290
	val provedgoal = Drule.close_derivation provedgoal
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   291
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   292
        val _ = print "Getting gif"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   293
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   294
        val graph_is_function = (provedgoal COMP conjunctionD1)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   295
                                  |> forall_elim_vars 0
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   296
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   297
        val _ = print "Getting cases"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   298
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   299
        val complete_thm = provedgoal COMP conjunctionD2
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   300
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   301
        val _ = print "making f_iff"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   302
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   303
	val f_iff = (graph_is_function RS ex1_iff) 
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   304
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   305
	val _ = Output.debug "Proving simplification rules"
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   306
	val psimps  = map2 (mk_psimp thy globals R f_iff graph_is_function) clauses values
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   307
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   308
	val _ = Output.debug "Proving partial induction rule"
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   309
	val (subset_pinduct, simple_pinduct) = mk_partial_induct_rule thy globals R complete_thm clauses
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   310
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   311
	val _ = Output.debug "Proving nested termination rule"
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   312
	val total_intro = mk_nest_term_rule thy globals R R_cases clauses
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   313
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   314
	val _ = Output.debug "Proving domain introduction rules"
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19930
diff changeset
   315
	val dom_intros = map (mk_domain_intro thy globals R R_cases) clauses
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   316
    in 
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   317
	FundefResult {f=f, G=G, R=R, completeness=complete_thm, 
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   318
	 psimps=psimps, subset_pinduct=subset_pinduct, simple_pinduct=simple_pinduct, total_intro=total_intro,
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   319
	 dom_intros=dom_intros}
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   320
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   321
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   322
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   323
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   324
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   325