TFL/post.sml
author nipkow
Thu, 30 Mar 2000 19:44:11 +0200
changeset 8622 870a58dd0ddd
parent 8526 0be2c98f15a7
child 8631 a10ae360b63c
permissions -rw-r--r--
the simplification rules returned from TFL are now paired with the row they came from.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     1
(*  Title:      TFL/post
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     2
    ID:         $Id$
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     4
    Copyright   1997  University of Cambridge
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     5
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     6
Postprocessing of TFL definitions
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     7
*)
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     8
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
     9
signature TFL = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    10
  sig
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    11
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    12
   val trace : bool ref
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    13
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
   structure Prim : TFL_sig
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    15
   val quiet_mode : bool ref
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    16
   val message : string -> unit
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    17
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    18
   val tgoalw : theory -> thm list -> thm list -> thm list
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    19
   val tgoal: theory -> thm list -> thm list
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    20
3405
2cccd0e3e9ea Numerous simplifications and removal of HOL-isms
paulson
parents: 3391
diff changeset
    21
   val std_postprocessor : simpset -> theory 
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    22
                           -> {induction:thm, rules:thm, TCs:term list list} 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    23
                           -> {induction:thm, rules:thm, nested_tcs:thm list}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    24
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    25
   val define_i : theory -> xstring -> term 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    26
                  -> simpset * thm list (*allows special simplication*)
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    27
                  -> term list
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
    28
                  -> theory * {rules:(thm*int)list, induct:thm, tcs:term list}
3331
c81c7f8ad333 Now checks the name of the function being defined
paulson
parents: 3302
diff changeset
    29
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    30
   val define   : theory -> xstring -> string 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    31
                  -> simpset * thm list 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    32
                  -> string list 
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
    33
                  -> theory * {rules:(thm*int)list, induct:thm, tcs:term list}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    34
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    35
   val defer_i : theory -> xstring
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    36
                  -> thm list (* congruence rules *)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    37
                  -> term list -> theory * thm
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    38
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
    39
   val defer : theory -> xstring
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    40
                  -> thm list 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    41
                  -> string list -> theory * thm
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
    42
(*
3459
112cbb8301dc Removal of structure Context and its replacement by a theorem list of
paulson
parents: 3405
diff changeset
    43
   val simplify_defn : simpset * thm list 
112cbb8301dc Removal of structure Context and its replacement by a theorem list of
paulson
parents: 3405
diff changeset
    44
                        -> theory * (string * Prim.pattern list)
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    45
                        -> {rules:thm list, induct:thm, tcs:term list}
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
    46
*)
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    47
  end;
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    48
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    49
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    50
structure Tfl: TFL =
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    51
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    52
 structure Prim = Prim
3391
5e45dd3b64e9 More de-HOLification: using Free, Const, etc. instead of mk_var, mk_const
paulson
parents: 3331
diff changeset
    53
 structure S = USyntax
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    54
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    55
 (* messages *)
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    56
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    57
 val quiet_mode = ref false;
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    58
 fun message s = if ! quiet_mode then () else writeln s;
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    59
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    60
 val trace = Prim.trace
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    61
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    62
 (*---------------------------------------------------------------------------
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    63
  * Extract termination goals so that they can be put it into a goalstack, or 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    64
  * have a tactic directly applied to them.
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    65
  *--------------------------------------------------------------------------*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    66
 fun termination_goals rules = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    67
     map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    68
       (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    69
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    70
 (*---------------------------------------------------------------------------
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    71
  * Finds the termination conditions in (highly massaged) definition and 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    72
  * puts them into a goalstack.
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    73
  *--------------------------------------------------------------------------*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    74
 fun tgoalw thy defs rules = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    75
   case termination_goals rules of
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    76
       [] => error "tgoalw: no termination conditions to prove"
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    77
     | L  => goalw_cterm defs 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    78
	       (Thm.cterm_of (Theory.sign_of thy) 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    79
			 (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    80
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    81
 fun tgoal thy = tgoalw thy [];
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    82
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    83
 (*---------------------------------------------------------------------------
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    84
 * Three postprocessors are applied to the definition.  It
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    85
 * attempts to prove wellfoundedness of the given relation, simplifies the
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    86
 * non-proved termination conditions, and finally attempts to prove the 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    87
 * simplified termination conditions.
3405
2cccd0e3e9ea Numerous simplifications and removal of HOL-isms
paulson
parents: 3391
diff changeset
    88
 *--------------------------------------------------------------------------*)
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    89
 fun std_postprocessor ss =
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    90
   Prim.postprocess
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    91
    {WFtac      = REPEAT (ares_tac [wf_empty, wf_pred_nat, 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    92
				    wf_measure, wf_inv_image, 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    93
				    wf_lex_prod, wf_less_than, wf_trancl] 1),
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    94
     terminator = asm_simp_tac ss 1
8526
0be2c98f15a7 replaced best_tac by force_tac, allowing some arithmetic reasoning
paulson
parents: 7262
diff changeset
    95
		  THEN TRY(CLASET' (fn cs => force_tac
0be2c98f15a7 replaced best_tac by force_tac, allowing some arithmetic reasoning
paulson
parents: 7262
diff changeset
    96
			   (cs addSDs [not0_implies_Suc], ss)) 1),
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    97
     simplifier = Rules.simpl_conv ss []};
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    98
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    99
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   100
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   101
 val concl = #2 o Rules.dest_thm;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   102
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   103
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   104
 * Postprocess a definition made by "define". This is a separate stage of 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   105
 * processing from the definition stage.
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   106
 *---------------------------------------------------------------------------*)
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   107
 local 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   108
 structure R = Rules
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   109
 structure U = Utils
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   110
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   111
 (* The rest of these local definitions are for the tricky nested case *)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   112
 val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   113
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   114
 fun id_thm th = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   115
    let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th))))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   116
    in lhs aconv rhs
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   117
    end handle _ => false
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   118
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   119
 fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   120
 val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   121
 val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   122
 fun mk_meta_eq r = case concl_of r of
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   123
      Const("==",_)$_$_ => r
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   124
   |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   125
   |   _ => r RS P_imp_P_eq_True
3405
2cccd0e3e9ea Numerous simplifications and removal of HOL-isms
paulson
parents: 3391
diff changeset
   126
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   127
 (*Is this the best way to invoke the simplifier??*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   128
 fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   129
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   130
 fun join_assums th = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   131
   let val {sign,...} = rep_thm th
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   132
       val tych = cterm_of sign
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   133
       val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   134
       val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   135
       val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   136
       val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   137
   in 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   138
     R.GEN_ALL 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   139
       (R.DISCH_ALL 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   140
	  (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   141
   end
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   142
   val gen_all = S.gen_all
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   143
 in
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   144
 fun proof_stage ss theory {f, R, rules, full_pats_TCs, TCs} =
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
   145
   let val dummy = message "Proving induction theorem ..."
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   146
       val ind = Prim.mk_induction theory 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   147
                    {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   148
       val dummy = (message "Proved induction theorem."; 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   149
		    message "Postprocessing ...");
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   150
       val {rules, induction, nested_tcs} = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   151
	   std_postprocessor ss theory {rules=rules, induction=ind, TCs=TCs}
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   152
   in
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   153
   case nested_tcs
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
   154
   of [] => (message "Postprocessing done.";
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   155
	     {induction=induction, rules=rules,tcs=[]})
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
   156
   | L  => let val dummy = message "Simplifying nested TCs ..."
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   157
	       val (solved,simplified,stubborn) =
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   158
		U.itlist (fn th => fn (So,Si,St) =>
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   159
		      if (id_thm th) then (So, Si, th::St) else
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   160
		      if (solved th) then (th::So, Si, St) 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   161
		      else (So, th::Si, St)) nested_tcs ([],[],[])
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   162
	       val simplified' = map join_assums simplified
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   163
	       val rewr = full_simplify (ss addsimps (solved @ simplified'));
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   164
	       val induction' = rewr induction
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   165
	       and rules'     = rewr rules
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
   166
	       val dummy = message "Postprocessing done."
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   167
	   in
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   168
	   {induction = induction',
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   169
		rules = rules',
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   170
		  tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   171
			    (simplified@stubborn)}
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   172
	   end
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   173
   end;
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   174
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   175
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   176
 (*lcp: curry the predicate of the induction rule*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   177
 fun curry_rule rl = split_rule_var
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   178
			 (head_of (HOLogic.dest_Trueprop (concl_of rl)), 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   179
			  rl);
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   180
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   181
 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   182
 val meta_outer = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   183
     curry_rule o standard o 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   184
     rule_by_tactic (REPEAT 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   185
		     (FIRSTGOAL (resolve_tac [allI, impI, conjI]
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   186
				 ORELSE' etac conjE)));
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   187
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   188
 (*Strip off the outer !P*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   189
 val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
3459
112cbb8301dc Removal of structure Context and its replacement by a theorem list of
paulson
parents: 3405
diff changeset
   190
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   191
 (*this function could be combined with define_i --lcp*)
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   192
 fun simplify_defn (ss, tflCongs) (thy,(id,pats)) =
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   193
    let val dummy = deny (id mem (Sign.stamp_names_of (Theory.sign_of thy)))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   194
			 ("Recursive definition " ^ id ^ 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   195
			  " would clash with the theory of the same name!")
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   196
	val def = freezeT(get_def thy id)   RS   meta_eq_to_obj_eq
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
   197
	val {theory,rules,rows,TCs,full_pats_TCs} = 
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   198
		 Prim.post_definition (Prim.congs tflCongs)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   199
		    (thy, (def,pats))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   200
	val {lhs=f,rhs} = S.dest_eq(concl def)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   201
	val (_,[R,_]) = S.strip_comb rhs
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   202
	val ss' = ss addsimps Prim.default_simps
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   203
	val {induction, rules, tcs} = 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   204
	      proof_stage ss' theory 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   205
		{f = f, R = R, rules = rules,
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   206
		 full_pats_TCs = full_pats_TCs,
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   207
		 TCs = TCs}
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   208
	val rules' = map (standard o normalize_thm [RSmp]) (R.CONJUNCTS rules)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   209
    in  {induct = meta_outer
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   210
		   (normalize_thm [RSspec,RSmp] (induction RS spec')), 
8622
870a58dd0ddd the simplification rules returned from TFL are now paired with the row they
nipkow
parents: 8526
diff changeset
   211
	 rules = ListPair.zip(rules', rows),
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   212
	 tcs = (termination_goals rules') @ tcs}
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   213
    end
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   214
   handle Utils.ERR {mesg,func,module} => 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   215
		error (mesg ^ 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   216
		       "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   217
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   218
(*---------------------------------------------------------------------------
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   219
 * Defining a function with an associated termination relation. 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   220
 *---------------------------------------------------------------------------*)
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   221
 fun define_i thy fid R ss_congs eqs = 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   222
   let val {functional,pats} = Prim.mk_functional thy eqs
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   223
       val thy = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   224
   in (thy, simplify_defn ss_congs (thy, (fid, pats)))
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   225
   end;
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   226
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   227
 fun define thy fid R ss_congs seqs = 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   228
   let val _ =  writeln ("Recursive function " ^ fid)
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   229
       val read = readtm (Theory.sign_of thy) (TVar(("DUMMY",0),[])) 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   230
   in  define_i thy fid (read R) ss_congs (map read seqs)  end
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   231
   handle Utils.ERR {mesg,...} => error mesg;
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   232
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   233
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   234
(*---------------------------------------------------------------------------
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   235
 *
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   236
 *     Definitions with synthesized termination relation
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   237
 *
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   238
 *---------------------------------------------------------------------------*)
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   239
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   240
 local open USyntax
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   241
 in 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   242
 fun func_of_cond_eqn tm =
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   243
   #1(strip_comb(#lhs(dest_eq(#2 (strip_forall(#2(strip_imp tm)))))))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   244
 end;
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   245
6554
5be3f13193d7 tuned defer_recdef interfaces;
wenzelm
parents: 6524
diff changeset
   246
 fun defer_i thy fid congs eqs = 
5be3f13193d7 tuned defer_recdef interfaces;
wenzelm
parents: 6524
diff changeset
   247
  let val {rules,R,theory,full_pats_TCs,SV,...} = 
5be3f13193d7 tuned defer_recdef interfaces;
wenzelm
parents: 6524
diff changeset
   248
	      Prim.lazyR_def thy (Sign.base_name fid) congs eqs
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   249
      val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules))
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   250
      val dummy = (message "Definition made."; 
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   251
		   message "Proving induction theorem ...");
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   252
      val induction = Prim.mk_induction theory 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   253
                         {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
   254
      val dummy = message "Induction theorem proved."
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   255
  in (theory, 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   256
      (*return the conjoined induction rule and recursion equations, 
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   257
	with assumptions remaining to discharge*)
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   258
      standard (induction RS (rules RS conjI)))
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   259
  end
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   260
6554
5be3f13193d7 tuned defer_recdef interfaces;
wenzelm
parents: 6524
diff changeset
   261
 fun defer thy fid congs seqs = 
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   262
   let val read = readtm (Theory.sign_of thy) (TVar(("DUMMY",0),[]))
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   263
   in  defer_i thy fid congs (map read seqs)  end
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   264
   handle Utils.ERR {mesg,...} => error mesg;
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   265
 end;
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   266
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   267
end;