| author | nipkow | 
| Thu, 12 Feb 2004 00:28:23 +0100 | |
| changeset 14385 | 6b15793a641a | 
| parent 13917 | a67c9e6570ac | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 11522 | 1 | (* Title: Pure/Proof/proof_rewrite_rules.ML | 
| 2 | ID: $Id$ | |
| 11539 | 3 | Author: Stefan Berghofer, TU Muenchen | 
| 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | |
| 11522 | 5 | |
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 6 | Simplification functions for proof terms involving meta level rules. | 
| 11522 | 7 | *) | 
| 8 | ||
| 9 | signature PROOF_REWRITE_RULES = | |
| 10 | sig | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 11 | val rew : bool -> typ list -> Proofterm.proof -> Proofterm.proof option | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 12 | val rprocs : bool -> (string * (typ list -> Proofterm.proof -> Proofterm.proof option)) list | 
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 13 | val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof | 
| 13341 | 14 | val elim_defs : Sign.sg -> bool -> thm list -> Proofterm.proof -> Proofterm.proof | 
| 13608 | 15 | val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof | 
| 12237 | 16 | val setup : (theory -> theory) list | 
| 11522 | 17 | end; | 
| 18 | ||
| 19 | structure ProofRewriteRules : PROOF_REWRITE_RULES = | |
| 20 | struct | |
| 21 | ||
| 22 | open Proofterm; | |
| 23 | ||
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 24 | fun rew b = | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 25 | let | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 26 | fun ? x = if b then Some x else None; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 27 | fun ax (prf as PAxm (s, prop, _)) Ts = | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 28 | if b then PAxm (s, prop, Some Ts) else prf; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 29 | fun ty T = if b then | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 30 | let val Type (_, [Type (_, [U, _]), _]) = T | 
| 13257 | 31 | in Some U end | 
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 32 | else None; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 33 | val equal_intr_axm = ax equal_intr_axm []; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 34 | val equal_elim_axm = ax equal_elim_axm []; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 35 | val symmetric_axm = ax symmetric_axm [propT]; | 
| 11522 | 36 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 37 |     fun rew' _ (PThm (("ProtoPure.rev_triv_goal", _), _, _, _) % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 38 |         (PThm (("ProtoPure.triv_goal", _), _, _, _) % _ %% prf)) = Some prf
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 39 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 40 |         (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = Some prf
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 41 |       | rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 42 |         (PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 43 | Some (equal_intr_axm % B % A %% prf2 %% prf1) | 
| 12002 | 44 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 45 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
 | 
| 12002 | 46 |         (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
 | 
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 47 |           _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 48 |         ((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 49 | Some (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) | 
| 12002 | 50 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 51 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 52 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 53 |           (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 54 |              _ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1)) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 55 |         ((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 56 | Some (tg %> B %% (equal_elim_axm %> A %> B %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 57 | (symmetric_axm % ? B % ? A %% prf1) %% prf2)) | 
| 11522 | 58 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 59 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
 | 
| 11612 | 60 |         (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 61 |           (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 62 |              (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 63 | let | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 64 | val _ $ A $ C = Envir.beta_norm X; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 65 | val _ $ B $ D = Envir.beta_norm Y | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 66 |         in Some (AbsP ("H1", ? X, AbsP ("H2", ? B,
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 67 | equal_elim_axm %> C %> D %% incr_pboundvars 2 0 prf2 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 68 | (PBound 1 %% (equal_elim_axm %> B %> A %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 69 | (symmetric_axm % ? A % ? B %% incr_pboundvars 2 0 prf1) %% PBound 0))))) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 70 | end | 
| 11522 | 71 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 72 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 73 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 74 |           (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 75 |             (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 76 |                (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 77 | let | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 78 | val _ $ A $ C = Envir.beta_norm Y; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 79 | val _ $ B $ D = Envir.beta_norm X | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 80 |         in Some (AbsP ("H1", ? X, AbsP ("H2", ? A,
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 81 | equal_elim_axm %> D %> C %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 82 | (symmetric_axm % ? C % ? D %% incr_pboundvars 2 0 prf2) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 83 | %% (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: 
12237diff
changeset | 84 | end | 
| 11522 | 85 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 86 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
 | 
| 11612 | 87 |         (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
 | 
| 88 |           (PAxm ("ProtoPure.reflexive", _, _) % _) %%
 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 89 |             (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 90 | let | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 91 | val Const (_, T) $ P = Envir.beta_norm X; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 92 | val _ $ Q = Envir.beta_norm Y; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 93 |         in Some (AbsP ("H", ? X, Abst ("x", ty T,
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 94 | 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: 
12237diff
changeset | 95 | (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: 
12237diff
changeset | 96 | end | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 97 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 98 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 99 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%        
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 100 |           (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 101 |             (PAxm ("ProtoPure.reflexive", _, _) % _) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 102 |               (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 103 | let | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 104 | val Const (_, T) $ P = Envir.beta_norm X; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 105 | val _ $ Q = Envir.beta_norm Y; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 106 | val t = incr_boundvars 1 P $ Bound 0; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 107 | val u = incr_boundvars 1 Q $ Bound 0 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 108 |         in Some (AbsP ("H", ? X, Abst ("x", ty T,
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 109 | equal_elim_axm %> t %> u %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 110 | (symmetric_axm % ? u % ? t %% (incr_pboundvars 1 1 prf %> Bound 0)) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 111 | %% (PBound 0 %> Bound 0)))) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 112 | end | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 113 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 114 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 115 |         (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2) %% prf3) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 116 | Some (equal_elim_axm %> B %> C %% prf2 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 117 | (equal_elim_axm %> A %> B %% prf1 %% prf3)) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 118 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 119 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 120 |           (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2)) %% prf3) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 121 | Some (equal_elim_axm %> B %> C %% (symmetric_axm % ? C % ? B %% prf1) %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 122 | (equal_elim_axm %> A %> B %% (symmetric_axm % ? B % ? A %% prf2) %% prf3)) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 123 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 124 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 125 |         (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = Some prf
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 126 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 127 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 128 |           (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = Some prf
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 129 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 130 |       | rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 131 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = Some prf
 | 
| 11522 | 132 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 133 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 134 |         (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 135 |           (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 136 |             (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 137 |               (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 138 | Some (equal_elim_axm %> C %> D %% prf2 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 139 | (equal_elim_axm %> A %> C %% prf3 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 140 | (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %% prf4))) | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 141 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 142 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 143 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 144 |           (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 145 |             (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 146 |               (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 147 |                 (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 148 | Some (equal_elim_axm %> A %> B %% prf1 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 149 | (equal_elim_axm %> C %> A %% (symmetric_axm % ? A % ? C %% prf3) %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 150 | (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %% prf4))) | 
| 11522 | 151 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 152 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 153 |         (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 154 |           (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 155 |             (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 156 |               (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 157 |                 (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 158 | Some (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 159 | (equal_elim_axm %> B %> D %% prf3 %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 160 | (equal_elim_axm %> A %> B %% prf1 %% prf4))) | 
| 11522 | 161 | |
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 162 |       | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 163 |         (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 164 |           (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 165 |             (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 166 |               (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 167 |                 (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 168 |                   (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 169 | Some (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 170 | (equal_elim_axm %> D %> B %% (symmetric_axm % ? B % ? D %% prf3) %% | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 171 | (equal_elim_axm %> C %> D %% prf2 %% prf4))) | 
| 11522 | 172 | |
| 13257 | 173 |       | rew' _ ((prf as PAxm ("ProtoPure.combination", _, _) %
 | 
| 174 |         Some ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
 | |
| 175 |           (PAxm ("ProtoPure.reflexive", _, _) % _)) =
 | |
| 176 | let val (U, V) = (case T of | |
| 177 | Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT)) | |
| 178 | in Some (prf %% (ax combination_axm [V, U] %> eq % ? eq % ? t % ? t %% | |
| 179 | (ax reflexive_axm [T] % ? eq) %% (ax reflexive_axm [U] % ? t))) | |
| 180 | end | |
| 181 | ||
| 12866 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 182 | | rew' _ _ = None; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 183 | in rew' end; | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 184 | |
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 185 | fun rprocs b = [("Pure/meta_equality", rew b)];
 | 
| 
c00df7765656
Rewrite procedure now works for both compact and full proof objects.
 berghofe parents: 
12237diff
changeset | 186 | val setup = [Proofterm.add_prf_rprocs (rprocs false)]; | 
| 11522 | 187 | |
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 188 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 189 | (**** apply rewriting function to all terms in proof ****) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 190 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 191 | fun rewrite_terms r = | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 192 | let | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 193 | fun rew_term Ts t = | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 194 | let | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 195 | val frees = map Free (variantlist | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 196 | (replicate (length Ts) "x", add_term_names (t, [])) ~~ Ts); | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 197 | val t' = r (subst_bounds (frees, t)); | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 198 | fun strip [] t = t | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 199 | | strip (_ :: xs) (Abs (_, _, t)) = strip xs t; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 200 | in | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 201 | strip Ts (foldl (uncurry lambda o Library.swap) (t', frees)) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 202 | end; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 203 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 204 | fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 205 | | rew Ts (prf % Some t) = rew Ts prf % Some (rew_term Ts t) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 206 | | rew Ts (Abst (s, Some T, prf)) = Abst (s, Some T, rew (T :: Ts) prf) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 207 | | rew Ts (AbsP (s, Some t, prf)) = AbsP (s, Some (rew_term Ts t), rew Ts prf) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 208 | | rew _ prf = prf | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 209 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 210 | in rew [] end; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 211 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 212 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 213 | (**** eliminate definitions in proof ****) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 214 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 215 | fun vars_of t = rev (foldl_aterms | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 216 | (fn (vs, v as Var _) => v ins vs | (vs, _) => vs) ([], t)); | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 217 | |
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 218 | fun insert_refl defs Ts (prf1 %% prf2) = | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 219 | insert_refl defs Ts prf1 %% insert_refl defs Ts prf2 | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 220 | | insert_refl defs Ts (Abst (s, Some T, prf)) = | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 221 | Abst (s, Some T, insert_refl defs (T :: Ts) prf) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 222 | | insert_refl defs Ts (AbsP (s, t, prf)) = | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 223 | AbsP (s, t, insert_refl defs Ts prf) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 224 | | insert_refl defs Ts prf = (case strip_combt prf of | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 225 | (PThm ((s, _), _, prop, Some Ts), ts) => | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 226 | if s mem defs then | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 227 | let | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 228 | val vs = vars_of prop; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 229 | val tvars = term_tvars prop; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 230 | val (_, rhs) = Logic.dest_equals prop; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 231 | val rhs' = foldl betapply (subst_TVars (map fst tvars ~~ Ts) | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 232 |                 (foldr (fn p => Abs ("", dummyT, abstract_over p)) (vs, rhs)),
 | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 233 | map the ts); | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 234 | in | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 235 | change_type (Some [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs' | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 236 | end | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 237 | else prf | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 238 | | (_, []) => prf | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 239 | | (prf', ts) => proof_combt' (insert_refl defs Ts prf', ts)); | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 240 | |
| 13341 | 241 | fun elim_defs sign r defs prf = | 
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 242 | let | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 243 | val tsig = Sign.tsig_of sign; | 
| 13341 | 244 | val defs' = map (Logic.dest_equals o prop_of o Drule.abs_def) defs | 
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 245 | val defnames = map Thm.name_of_thm defs; | 
| 13341 | 246 | val f = if not r then I else | 
| 247 | let | |
| 248 | val cnames = map (fst o dest_Const o fst) defs'; | |
| 249 | val thms = flat (map (fn (s, ps) => | |
| 250 | if s mem defnames then [] | |
| 251 | else map (pair s o Some o fst) (filter_out (fn (p, _) => | |
| 13646 | 252 | null (term_consts p inter cnames)) ps)) | 
| 13341 | 253 | (Symtab.dest (thms_of_proof Symtab.empty prf))) | 
| 254 | in Reconstruct.expand_proof sign thms end | |
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 255 | in | 
| 13341 | 256 | rewrite_terms (Pattern.rewrite_term tsig defs' []) | 
| 257 | (insert_refl defnames [] (f prf)) | |
| 12906 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 258 | end; | 
| 
165f4e1937f4
New function for eliminating definitions in proof term.
 berghofe parents: 
12866diff
changeset | 259 | |
| 13608 | 260 | |
| 261 | (**** eliminate all variables that don't occur in the proposition ****) | |
| 262 | ||
| 263 | fun elim_vars mk_default prf = | |
| 264 | let | |
| 265 | val prop = Reconstruct.prop_of prf; | |
| 13917 | 266 | val tv = term_vars prop; | 
| 267 | val tf = term_frees prop; | |
| 268 | ||
| 269 | fun mk_default' T = list_abs | |
| 270 | (apfst (map (pair "x")) (apsnd mk_default (strip_type T))); | |
| 271 | ||
| 272 | fun elim_varst (t $ u) = elim_varst t $ elim_varst u | |
| 273 | | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t) | |
| 274 | | elim_varst (f as Free (_, T)) = if f mem tf then f else mk_default' T | |
| 275 | | elim_varst (v as Var (_, T)) = if v mem tv then v else mk_default' T | |
| 276 | | elim_varst t = t | |
| 13608 | 277 | in | 
| 13917 | 278 | map_proof_terms (fn t => if not (null (term_vars t \\ tv)) orelse | 
| 279 | not (null (term_frees t \\ tf)) then Envir.beta_norm (elim_varst t) | |
| 280 | else t) I prf | |
| 13608 | 281 | end; | 
| 282 | ||
| 11522 | 283 | end; |