src/Pure/Proof/proof_rewrite_rules.ML
author paulson
Fri, 05 Oct 2007 09:59:03 +0200
changeset 24854 0ebcd575d3c6
parent 23178 07ba6b58b3d2
child 26424 a6cad32a27b0
permissions -rw-r--r--
filtering out some package theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     1
(*  Title:      Pure/Proof/proof_rewrite_rules.ML
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     4
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
     5
Simplification functions for proof terms involving meta level rules.
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     6
*)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     7
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
signature PROOF_REWRITE_RULES =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     9
sig
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    10
  val rew : bool -> typ list -> Proofterm.proof -> Proofterm.proof option
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    11
  val rprocs : bool -> (string * (typ list -> Proofterm.proof -> Proofterm.proof option)) list
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
    12
  val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17137
diff changeset
    13
  val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
    14
  val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof
22280
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
    15
  val hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
    16
  val un_hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    17
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    18
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    19
structure ProofRewriteRules : PROOF_REWRITE_RULES =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    20
struct
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    21
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
open Proofterm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    23
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    24
fun rew b _ =
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    25
  let
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    26
    fun ?? x = if b then SOME x else NONE;
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    27
    fun ax (prf as PAxm (s, prop, _)) Ts =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    28
      if b then PAxm (s, prop, SOME Ts) else prf;
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    29
    fun ty T = if b then
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    30
        let val Type (_, [Type (_, [U, _]), _]) = T
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    31
        in SOME U end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    32
      else NONE;
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    33
    val equal_intr_axm = ax equal_intr_axm [];
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    34
    val equal_elim_axm = ax equal_elim_axm [];
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    35
    val symmetric_axm = ax symmetric_axm [propT];
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    36
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    37
    fun rew' (PThm ("ProtoPure.protectD", _, _, _) % _ %%
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    38
        (PThm ("ProtoPure.protectI", _, _, _) % _ %% prf)) = SOME prf
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    39
      | rew' (PThm ("ProtoPure.conjunctionD1", _, _, _) % _ % _ %%
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    40
        (PThm ("ProtoPure.conjunctionI", _, _, _) % _ % _ %% prf %% _)) = SOME prf
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    41
      | rew' (PThm ("ProtoPure.conjunctionD2", _, _, _) % _ % _ %%
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    42
        (PThm ("ProtoPure.conjunctionI", _, _, _) % _ % _ %% _ %% prf)) = SOME prf
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    43
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    44
        (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    45
      | rew' (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    46
        (PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    47
            SOME (equal_intr_axm % B % A %% prf2 %% prf1)
12002
bc9b5bad0e7b Additional rules for simplifying inside "Goal"
berghofe
parents: 11612
diff changeset
    48
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    49
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
18024
853e8219732a renamed Goal constant to prop;
wenzelm
parents: 17979
diff changeset
    50
        (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("prop", _)) %
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    51
          _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %%
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    52
        ((tg as PThm ("ProtoPure.protectI", _, _, _)) % _ %% prf2)) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    53
        SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))
12002
bc9b5bad0e7b Additional rules for simplifying inside "Goal"
berghofe
parents: 11612
diff changeset
    54
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    55
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    56
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
18024
853e8219732a renamed Goal constant to prop;
wenzelm
parents: 17979
diff changeset
    57
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("prop", _)) %
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    58
             _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1)) %%
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
    59
        ((tg as PThm ("ProtoPure.protectI", _, _, _)) % _ %% prf2)) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    60
        SOME (tg %> B %% (equal_elim_axm %> A %> B %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    61
          (symmetric_axm % ?? B % ?? A %% prf1) %% prf2))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    62
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    63
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    64
        (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    65
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    66
             (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    67
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    68
          val _ $ A $ C = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    69
          val _ $ B $ D = Envir.beta_norm Y
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    70
        in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B,
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    71
          equal_elim_axm %> C %> D %% incr_pboundvars 2 0 prf2 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    72
            (PBound 1 %% (equal_elim_axm %> B %> A %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    73
              (symmetric_axm % ?? A % ?? B %% incr_pboundvars 2 0 prf1) %% PBound 0)))))
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    74
        end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    75
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    76
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    77
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    78
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    79
            (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    80
               (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    81
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    82
          val _ $ A $ C = Envir.beta_norm Y;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    83
          val _ $ B $ D = Envir.beta_norm X
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    84
        in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? A,
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    85
          equal_elim_axm %> D %> C %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    86
            (symmetric_axm % ?? C % ?? D %% incr_pboundvars 2 0 prf2)
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    87
              %% (PBound 1 %% (equal_elim_axm %> A %> B %% incr_pboundvars 2 0 prf1 %% PBound 0)))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    88
        end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    89
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
    90
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
    91
        (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    92
          (PAxm ("ProtoPure.reflexive", _, _) % _) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    93
            (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    94
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    95
          val Const (_, T) $ P = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    96
          val _ $ Q = Envir.beta_norm Y;
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
    97
        in SOME (AbsP ("H", ?? X, Abst ("x", ty T,
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    98
            equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
    99
              (incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   100
        end
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   101
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   102
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   103
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%        
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   104
          (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   105
            (PAxm ("ProtoPure.reflexive", _, _) % _) %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   106
              (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   107
        let
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   108
          val Const (_, T) $ P = Envir.beta_norm X;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   109
          val _ $ Q = Envir.beta_norm Y;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   110
          val t = incr_boundvars 1 P $ Bound 0;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   111
          val u = incr_boundvars 1 Q $ Bound 0
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   112
        in SOME (AbsP ("H", ?? X, Abst ("x", ty T,
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   113
          equal_elim_axm %> t %> u %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   114
            (symmetric_axm % ?? u % ?? t %% (incr_pboundvars 1 1 prf %> Bound 0))
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   115
              %% (PBound 0 %> Bound 0))))
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   116
        end
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   117
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   118
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   119
        (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) =
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   120
           SOME (equal_elim_axm %> B %> C %% prf2 %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   121
             (equal_elim_axm %> A %> B %% prf1 %% prf3))
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   122
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   123
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   124
          (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) =
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   125
           SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %%
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   126
             (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf2) %% prf3))
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   127
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   128
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   129
        (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = SOME prf
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   130
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   131
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   132
          (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = SOME prf
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   133
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   134
      | rew' (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   135
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = SOME prf
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   136
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   137
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   138
        (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   139
          (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   140
            (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   141
              (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   142
          SOME (equal_elim_axm %> C %> D %% prf2 %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   143
            (equal_elim_axm %> A %> C %% prf3 %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   144
              (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% prf4)))
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   145
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   146
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   147
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   148
          (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   149
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   150
              (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   151
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   152
          SOME (equal_elim_axm %> A %> B %% prf1 %%
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   153
            (equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %%
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   154
              (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   155
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   156
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   157
        (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   158
          (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   159
            (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   160
              (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   161
                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   162
          SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   163
            (equal_elim_axm %> B %> D %% prf3 %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   164
              (equal_elim_axm %> A %> B %% prf1 %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   165
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   166
      | rew' (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   167
        (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   168
          (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   169
            (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   170
              (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   171
                (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   172
                  (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   173
          SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %%
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   174
            (equal_elim_axm %> D %> B %% (symmetric_axm % ?? B % ?? D %% prf3) %%
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   175
              (equal_elim_axm %> C %> D %% prf2 %% prf4)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   176
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   177
      | rew' ((prf as PAxm ("ProtoPure.combination", _, _) %
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   178
        SOME ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
13257
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   179
          (PAxm ("ProtoPure.reflexive", _, _) % _)) =
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   180
        let val (U, V) = (case T of
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   181
          Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT))
17137
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   182
        in SOME (prf %% (ax combination_axm [V, U] %> eq % ?? eq % ?? t % ?? t %%
0f48fbb60a61 replaced ? by ??
haftmann
parents: 17018
diff changeset
   183
          (ax reflexive_axm [T] % ?? eq) %% (ax reflexive_axm [U] % ?? t)))
13257
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   184
        end
1b7104a1c0bd Additional rule for rewriting on ==.
berghofe
parents: 13198
diff changeset
   185
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   186
      | rew' _ = NONE;
12866
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   187
  in rew' end;
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   188
c00df7765656 Rewrite procedure now works for both compact and full proof objects.
berghofe
parents: 12237
diff changeset
   189
fun rprocs b = [("Pure/meta_equality", rew b)];
22662
3e492ba59355 canonical merge operations
haftmann
parents: 22280
diff changeset
   190
val _ = Context.add_setup (fold Proofterm.add_prf_rproc (rprocs false));
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   191
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   192
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   193
(**** apply rewriting function to all terms in proof ****)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   194
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   195
fun rewrite_terms r =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   196
  let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   197
    fun rew_term Ts t =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   198
      let
20076
def4ad161528 Name.invent_list;
wenzelm
parents: 19482
diff changeset
   199
        val frees = map Free (Name.invent_list (add_term_names (t, [])) "xa" (length Ts) ~~ Ts);
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   200
        val t' = r (subst_bounds (frees, t));
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   201
        fun strip [] t = t
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   202
          | strip (_ :: xs) (Abs (_, _, t)) = strip xs t;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   203
      in
19473
wenzelm
parents: 19466
diff changeset
   204
        strip Ts (fold lambda frees t')
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   205
      end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   206
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   207
    fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   208
      | rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   209
      | rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   210
      | rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf)
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   211
      | rew _ prf = prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   212
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   213
  in rew [] end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   214
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   215
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   216
(**** eliminate definitions in proof ****)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   217
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16787
diff changeset
   218
fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []);
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   219
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   220
fun insert_refl defs Ts (prf1 %% prf2) =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   221
      insert_refl defs Ts prf1 %% insert_refl defs Ts prf2
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   222
  | insert_refl defs Ts (Abst (s, SOME T, prf)) =
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   223
      Abst (s, SOME T, insert_refl defs (T :: Ts) prf)
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   224
  | insert_refl defs Ts (AbsP (s, t, prf)) =
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   225
      AbsP (s, t, insert_refl defs Ts prf)
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   226
  | insert_refl defs Ts prf = (case strip_combt prf of
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
   227
        (PThm (s, _, prop, SOME Ts), ts) =>
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20076
diff changeset
   228
          if member (op =) defs s then
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   229
            let
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   230
              val vs = vars_of prop;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   231
              val tvars = term_tvars prop;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   232
              val (_, rhs) = Logic.dest_equals prop;
18185
9d51fad6bb1f Term.betapplys;
wenzelm
parents: 18024
diff changeset
   233
              val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts)
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22662
diff changeset
   234
                (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs),
19466
wenzelm
parents: 19309
diff changeset
   235
                map the ts);
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   236
            in
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   237
              change_type (SOME [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs'
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   238
            end
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   239
          else prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   240
      | (_, []) => prf
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   241
      | (prf', ts) => proof_combt' (insert_refl defs Ts prf', ts));
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   242
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17137
diff changeset
   243
fun elim_defs thy r defs prf =
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   244
  let
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   245
    val defs' = map (Logic.dest_equals o prop_of o Drule.abs_def) defs
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20664
diff changeset
   246
    val defnames = map Thm.get_name defs;
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   247
    val f = if not r then I else
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   248
      let
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   249
        val cnames = map (fst o dest_Const o fst) defs';
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19473
diff changeset
   250
        val thms = maps (fn (s, ps) =>
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20076
diff changeset
   251
            if member (op =) defnames s then []
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   252
            else map (pair s o SOME o fst) (filter_out (fn (p, _) =>
13646
46ed3d042ba5 Ported find_intro/elim to Isar.
nipkow
parents: 13608
diff changeset
   253
              null (term_consts p inter cnames)) ps))
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19473
diff changeset
   254
          (Symtab.dest (thms_of_proof prf Symtab.empty))
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17137
diff changeset
   255
      in Reconstruct.expand_proof thy thms end
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   256
  in
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17137
diff changeset
   257
    rewrite_terms (Pattern.rewrite_term thy defs' [])
13341
f15ed50d16cf - Moved abs_def to drule.ML
berghofe
parents: 13257
diff changeset
   258
      (insert_refl defnames [] (f prf))
12906
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   259
  end;
165f4e1937f4 New function for eliminating definitions in proof term.
berghofe
parents: 12866
diff changeset
   260
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   261
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   262
(**** eliminate all variables that don't occur in the proposition ****)
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   263
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   264
fun elim_vars mk_default prf =
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   265
  let
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   266
    val prop = Reconstruct.prop_of prf;
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   267
    val tv = Term.add_vars prop [];
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   268
    val tf = Term.add_frees prop [];
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   269
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   270
    fun hidden_variable (Var v) = not (member (op =) tv v)
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   271
      | hidden_variable (Free f) = not (member (op =) tf f)
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   272
      | hidden_variable _ = false;
13917
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   273
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   274
    fun mk_default' T = list_abs
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   275
      (apfst (map (pair "x")) (apsnd mk_default (strip_type T)));
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   276
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   277
    fun elim_varst (t $ u) = elim_varst t $ elim_varst u
a67c9e6570ac elim_vars now handles both Vars and Frees.
berghofe
parents: 13646
diff changeset
   278
      | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t)
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   279
      | elim_varst (t as Free (x, T)) = if member (op =) tf (x, T) then t else mk_default' T
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   280
      | elim_varst (t as Var (xi, T)) = if member (op =) tv (xi, T) then t else mk_default' T
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   281
      | elim_varst t = t;
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   282
  in
19309
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   283
    map_proof_terms (fn t =>
8ea43e9ad83a avoid polymorphic equality;
wenzelm
parents: 19126
diff changeset
   284
      if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf
13608
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   285
  end;
9a6f43b8eae1 Added function elim_vars.
berghofe
parents: 13341
diff changeset
   286
22280
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   287
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   288
(**** convert between hhf and non-hhf form ****)
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   289
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   290
fun hhf_proof P Q prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   291
  let
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   292
    val params = Logic.strip_params Q;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   293
    val Hs = Logic.strip_assums_hyp P;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   294
    val Hs' = Logic.strip_assums_hyp Q;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   295
    val k = length Hs;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   296
    val l = length params;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   297
    fun mk_prf i j Hs Hs' (Const ("all", _) $ Abs (_, _, P)) prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   298
          mk_prf i (j - 1) Hs Hs' P (prf %> Bound j)
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   299
      | mk_prf i j (H :: Hs) (H' :: Hs') (Const ("==>", _) $ _ $ P) prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   300
          mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i))
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   301
      | mk_prf _ _ _ _ _ prf = prf
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   302
  in
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   303
    prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |>
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   304
    fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |>
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   305
    fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   306
  end
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   307
and un_hhf_proof P Q prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   308
  let
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   309
    val params = Logic.strip_params Q;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   310
    val Hs = Logic.strip_assums_hyp P;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   311
    val Hs' = Logic.strip_assums_hyp Q;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   312
    val k = length Hs;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   313
    val l = length params;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   314
    fun mk_prf (Const ("all", _) $ Abs (s, T, P)) prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   315
          Abst (s, SOME T, mk_prf P prf)
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   316
      | mk_prf (Const ("==>", _) $ P $ Q) prf =
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   317
          AbsP ("H", SOME P, mk_prf Q prf)
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   318
      | mk_prf _ prf = prf
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   319
  in
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   320
    prf |> Proofterm.incr_pboundvars k l |>
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   321
    fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |>
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   322
    fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i))
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   323
      (Hs ~~ Hs' ~~ (k - 1 downto 0)) |>
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   324
    mk_prf Q
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   325
  end;
a20a203c8f41 Added functions hhf_proof and un_hhf_proof.
berghofe
parents: 21646
diff changeset
   326
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   327
end;