Exchanged % and %%.
authorberghofe
Fri, 28 Sep 2001 11:04:44 +0200
changeset 11612 ae8450657bf0
parent 11611 b0c69f4db64c
child 11613 c92a99a2b5b1
Exchanged % and %%.
src/Pure/Proof/proof_rewrite_rules.ML
src/Pure/Proof/proofchecker.ML
src/Pure/Thy/thm_deps.ML
src/Pure/thm.ML
--- a/src/Pure/Proof/proof_rewrite_rules.ML	Thu Sep 27 22:30:09 2001 +0200
+++ b/src/Pure/Proof/proof_rewrite_rules.ML	Fri Sep 28 11:04:44 2001 +0200
@@ -17,82 +17,82 @@
 
 open Proofterm;
 
-fun rew _ (PThm (("ProtoPure.rev_triv_goal", _), _, _, _) %% _ %
-      (PThm (("ProtoPure.triv_goal", _), _, _, _) %% _ % prf)) = Some prf
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% _ %% _ %
-      (PAxm ("ProtoPure.equal_intr", _, _) %% _ %% _ % prf % _)) = Some prf
-  | rew _ (PAxm ("ProtoPure.symmetric", _, _) %% _ %% _ %
-      (PAxm ("ProtoPure.equal_intr", _, _) %% A %% B % prf1 % prf2)) =
-          Some (equal_intr_axm %% B %% A % prf2 % prf1)
+fun rew _ (PThm (("ProtoPure.rev_triv_goal", _), _, _, _) % _ %%
+      (PThm (("ProtoPure.triv_goal", _), _, _, _) % _ %% prf)) = Some prf
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
+      (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = Some prf
+  | rew _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
+      (PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
+          Some (equal_intr_axm % B % A %% prf2 %% prf1)
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some X %% Some Y %
-      (PAxm ("ProtoPure.combination", _, _) %% _ %% _ %% _ %% _ %
-        (PAxm ("ProtoPure.combination", _, _) %% Some (Const ("==>", _)) %% _ %% _ %% _ %
-           (PAxm ("ProtoPure.reflexive", _, _) %% _) % prf1) % prf2)) =
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+      (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
+        (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
+           (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
       let
         val _ $ A $ C = Envir.beta_norm X;
         val _ $ B $ D = Envir.beta_norm Y
       in Some (AbsP ("H1", None, AbsP ("H2", None,
-        equal_elim_axm %%% C %%% D % incr_pboundvars 2 0 prf2 %
-          (PBound 1 % (equal_elim_axm %%% B %%% A %
-            (symmetric_axm %% None %% None % incr_pboundvars 2 0 prf1) % PBound 0)))))
+        equal_elim_axm %> C %> D %% incr_pboundvars 2 0 prf2 %%
+          (PBound 1 %% (equal_elim_axm %> B %> A %%
+            (symmetric_axm % None % None %% incr_pboundvars 2 0 prf1) %% PBound 0)))))
       end
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some X %% Some Y %
-      (PAxm ("ProtoPure.symmetric", _, _) %% _ %% _ %
-        (PAxm ("ProtoPure.combination", _, _) %% _ %% _ %% _ %% _ %
-          (PAxm ("ProtoPure.combination", _, _) %% Some (Const ("==>", _)) %% _ %% _ %% _ %
-             (PAxm ("ProtoPure.reflexive", _, _) %% _) % prf1) % prf2))) =
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+      (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
+        (PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
+          (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
+             (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
       let
         val _ $ A $ C = Envir.beta_norm Y;
         val _ $ B $ D = Envir.beta_norm X
       in Some (AbsP ("H1", None, AbsP ("H2", None,
-        equal_elim_axm %%% D %%% C %
-          (symmetric_axm %% None %% None % incr_pboundvars 2 0 prf2)
-            % (PBound 1 % (equal_elim_axm %%% A %%% B % incr_pboundvars 2 0 prf1 % PBound 0)))))
+        equal_elim_axm %> D %> C %%
+          (symmetric_axm % None % None %% incr_pboundvars 2 0 prf2)
+            %% (PBound 1 %% (equal_elim_axm %> A %> B %% incr_pboundvars 2 0 prf1 %% PBound 0)))))
       end
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some X %% Some Y %
-      (PAxm ("ProtoPure.combination", _, _) %% Some (Const ("all", _)) %% _ %% _ %% _ %
-        (PAxm ("ProtoPure.reflexive", _, _) %% _) %
-          (PAxm ("ProtoPure.abstract_rule", _, _) %% _ %% _ % prf))) =
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+      (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
+        (PAxm ("ProtoPure.reflexive", _, _) % _) %%
+          (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
       let
         val _ $ P = Envir.beta_norm X;
         val _ $ Q = Envir.beta_norm Y;
       in Some (AbsP ("H", None, Abst ("x", None,
-          equal_elim_axm %%% incr_boundvars 1 P $ Bound 0 %%% incr_boundvars 1 Q $ Bound 0 %
-            (incr_pboundvars 1 1 prf %%% Bound 0) % (PBound 0 %%% Bound 0))))
+          equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%
+            (incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0))))
       end
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some X %% Some Y %
-      (PAxm ("ProtoPure.symmetric", _, _) %% _ %% _ %        
-        (PAxm ("ProtoPure.combination", _, _) %% Some (Const ("all", _)) %% _ %% _ %% _ %
-          (PAxm ("ProtoPure.reflexive", _, _) %% _) %
-            (PAxm ("ProtoPure.abstract_rule", _, _) %% _ %% _ % prf)))) =
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+      (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%        
+        (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
+          (PAxm ("ProtoPure.reflexive", _, _) % _) %%
+            (PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
       let
         val _ $ P = Envir.beta_norm X;
         val _ $ Q = Envir.beta_norm Y;
       in Some (AbsP ("H", None, Abst ("x", None,
-        equal_elim_axm %%% incr_boundvars 1 P $ Bound 0 %%% incr_boundvars 1 Q $ Bound 0 %
-          (symmetric_axm %% None %% None % (incr_pboundvars 1 1 prf %%% Bound 0))
-            % (PBound 0 %%% Bound 0))))
+        equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%
+          (symmetric_axm % None % None %% (incr_pboundvars 1 1 prf %> Bound 0))
+            %% (PBound 0 %> Bound 0))))
       end
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some A %% Some C %
-      (PAxm ("ProtoPure.transitive", _, _) %% _ %% Some B %% _ % prf1 % prf2) % prf3) =
-         Some (equal_elim_axm %%% B %%% C % prf2 %
-           (equal_elim_axm %%% A %%% B % prf1 % prf3))
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% Some A %% Some C %
-      (PAxm ("ProtoPure.symmetric", _, _) %% _ %% _ %
-        (PAxm ("ProtoPure.transitive", _, _) %% _ %% Some B %% _ % prf1 % prf2)) % prf3) =
-         Some (equal_elim_axm %%% B %%% C % (symmetric_axm %% None %% None % prf1) %
-           (equal_elim_axm %%% A %%% B % (symmetric_axm %% None %% None % prf2) % prf3))
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
+      (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2) %% prf3) =
+         Some (equal_elim_axm %> B %> C %% prf2 %%
+           (equal_elim_axm %> A %> B %% prf1 %% prf3))
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
+      (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
+        (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2)) %% prf3) =
+         Some (equal_elim_axm %> B %> C %% (symmetric_axm % None % None %% prf1) %%
+           (equal_elim_axm %> A %> B %% (symmetric_axm % None % None %% prf2) %% prf3))
 
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% _ %% _ %
-      (PAxm ("ProtoPure.reflexive", _, _) %% _) % prf) = Some prf
-  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) %% _ %% _ %
-      (PAxm ("ProtoPure.symmetric", _, _) %% _ %% _ %
-        (PAxm ("ProtoPure.reflexive", _, _) %% _)) % prf) = Some prf
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
+      (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = Some prf
+  | rew _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
+      (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
+        (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = Some prf
 
   | rew _ _ = None;
 
--- a/src/Pure/Proof/proofchecker.ML	Thu Sep 27 22:30:09 2001 +0200
+++ b/src/Pure/Proof/proofchecker.ML	Fri Sep 28 11:04:44 2001 +0200
@@ -12,7 +12,7 @@
   val thm_of_proof : theory -> Proofterm.proof -> thm
 end;
 
-structure ProofChecker =
+structure ProofChecker : PROOF_CHECKER =
 struct
 
 open Proofterm;
@@ -72,7 +72,7 @@
             Thm.forall_intr (Thm.cterm_of sg (Free (x, T))) thm
           end
 
-      | thm_of vs Hs (prf %% Some t) =
+      | thm_of vs Hs (prf % Some t) =
           let
             val thm = thm_of vs Hs prf
             val ct = Thm.cterm_of sg (Term.subst_bounds (map Free vs, t))
@@ -86,7 +86,7 @@
             Thm.implies_intr ct thm
           end
 
-      | thm_of vs Hs (prf % prf') =
+      | thm_of vs Hs (prf %% prf') =
           let 
             val thm = beta_eta_convert (thm_of vs Hs prf);
             val thm' = beta_eta_convert (thm_of vs Hs prf')
--- a/src/Pure/Thy/thm_deps.ML	Thu Sep 27 22:30:09 2001 +0200
+++ b/src/Pure/Thy/thm_deps.ML	Fri Sep 28 11:04:44 2001 +0200
@@ -31,9 +31,9 @@
 
 fun make_deps_graph (p, AbsP (_, _, prf)) = make_deps_graph (p, prf)
   | make_deps_graph (p, Abst (_, _, prf)) = make_deps_graph (p, prf)
-  | make_deps_graph (p, prf1 % prf2) =
+  | make_deps_graph (p, prf1 %% prf2) =
       make_deps_graph (make_deps_graph (p, prf1), prf2)
-  | make_deps_graph (p, prf %% _) = make_deps_graph (p, prf)
+  | make_deps_graph (p, prf % _) = make_deps_graph (p, prf)
   | make_deps_graph (p, MinProof prfs) = foldl make_deps_graph (p, prfs)
   | make_deps_graph (p as (gra, parents), prf) =
       let val ((name, tags), prf') = dest_thm_axm prf
--- a/src/Pure/thm.ML	Thu Sep 27 22:30:09 2001 +0200
+++ b/src/Pure/thm.ML	Fri Sep 28 11:04:44 2001 +0200
@@ -590,7 +590,7 @@
                 if imp=implies andalso  A aconv propA
                 then
                   Thm{sign_ref= merge_thm_sgs(thAB,thA),
-                      der = Pt.infer_derivs (curry Pt.%) der derA,
+                      der = Pt.infer_derivs (curry Pt.%%) der derA,
                       maxidx = Int.max(maxA,maxidx),
                       shyps = union_sort (shypsA, shyps),
                       hyps = union_term(hypsA,hyps),  (*dups suppressed*)
@@ -635,7 +635,7 @@
               raise THM("forall_elim: type mismatch", 0, [th])
           else let val thm = fix_shyps [th] []
                     (Thm{sign_ref= Sign.merge_refs(sign_ref,sign_reft),
-                         der = Pt.infer_derivs' (Pt.%% o rpair (Some t)) der,
+                         der = Pt.infer_derivs' (Pt.% o rpair (Some t)) der,
                          maxidx = Int.max(maxidx, maxt),
                          shyps = [],
                          hyps = hyps,