diff -r c22e092b39e9 -r c0ab0824ccb5 src/HOL/ex/veriT_Preprocessing.thy --- a/src/HOL/ex/veriT_Preprocessing.thy Thu Feb 09 23:00:56 2017 +0100 +++ b/src/HOL/ex/veriT_Preprocessing.thy Fri Feb 10 08:27:27 2017 +0100 @@ -78,7 +78,7 @@ | str_of_rule_name (Let ts) = "Let[" ^ commas (map (Syntax.string_of_term @{context}) ts) ^ "]"; -datatype rule = Rule of rule_name * rule list; +datatype node = Node of rule_name * node list; fun lambda_count (Abs (_, _, t)) = lambda_count t + 1 | lambda_count ((t as Abs _) $ _) = lambda_count t - 1 @@ -162,7 +162,7 @@ | _ => raise TERM ("apply_Let_right", [lhs, rhs])) end; -fun reconstruct_proof ctxt (lrhs as (_, rhs), Rule (rule_name, prems)) = +fun reconstruct_proof ctxt (lrhs as (_, rhs), Node (rule_name, prems)) = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq lrhs); val ary = length prems; @@ -221,7 +221,7 @@ val proof0 = ((@{term "\x :: nat. p x"}, @{term "p (SOME x :: nat. p x)"}), - Rule (Sko_Ex, [Rule (Refl, [])])); + Node (Sko_Ex, [Node (Refl, [])])); reconstruct_proof @{context} proof0; \ @@ -230,7 +230,7 @@ val proof1 = ((@{term "\ (\x :: nat. \y :: nat. p x y)"}, @{term "\ (\y :: nat. p (SOME x :: nat. \ (\y :: nat. p x y)) y)"}), - Rule (Cong, [Rule (Sko_All, [Rule (Bind, [Rule (Refl, [])])])])); + Node (Cong, [Node (Sko_All, [Node (Bind, [Node (Refl, [])])])])); reconstruct_proof @{context} proof1; \ @@ -240,7 +240,7 @@ ((@{term "\x :: nat. \y :: nat. \z :: nat. p x y z"}, @{term "\x :: nat. p x (SOME y :: nat. \z :: nat. p x y z) (SOME z :: nat. p x (SOME y :: nat. \z :: nat. p x y z) z)"}), - Rule (Bind, [Rule (Sko_Ex, [Rule (Sko_Ex, [Rule (Refl, [])])])])); + Node (Bind, [Node (Sko_Ex, [Node (Sko_Ex, [Node (Refl, [])])])])); reconstruct_proof @{context} proof2 \ @@ -249,7 +249,7 @@ val proof3 = ((@{term "\x :: nat. \x :: nat. \x :: nat. p x x x"}, @{term "\x :: nat. p (SOME x :: nat. p x x x) (SOME x. p x x x) (SOME x. p x x x)"}), - Rule (Bind, [Rule (Sko_Ex, [Rule (Sko_Ex, [Rule (Refl, [])])])])); + Node (Bind, [Node (Sko_Ex, [Node (Sko_Ex, [Node (Refl, [])])])])); reconstruct_proof @{context} proof3 \ @@ -258,7 +258,7 @@ val proof4 = ((@{term "\x :: nat. \x :: nat. \x :: nat. p x x x"}, @{term "\x :: nat. \x :: nat. p (SOME x :: nat. p x x x) (SOME x. p x x x) (SOME x. p x x x)"}), - Rule (Bind, [Rule (Bind, [Rule (Sko_Ex, [Rule (Refl, [])])])])); + Node (Bind, [Node (Bind, [Node (Sko_Ex, [Node (Refl, [])])])])); reconstruct_proof @{context} proof4 \ @@ -268,19 +268,19 @@ ((@{term "\x :: nat. q \ (\x :: nat. \x :: nat. p x x x)"}, @{term "\x :: nat. q \ (\x :: nat. p (SOME x :: nat. p x x x) (SOME x. p x x x) (SOME x. p x x x))"}), - Rule (Bind, [Rule (Cong, [Rule (Refl, []), Rule (Bind, [Rule (Sko_Ex, - [Rule (Refl, [])])])])])); + Node (Bind, [Node (Cong, [Node (Refl, []), Node (Bind, [Node (Sko_Ex, + [Node (Refl, [])])])])])); reconstruct_proof @{context} proof5 \ ML \ val proof6 = - ((@{term "\ (\x :: nat. q \ (\x :: nat. \x :: nat. p x x x))"}, - @{term "\ (\x :: nat. q \ - (\x :: nat. p (SOME x :: nat. \ p x x x) (SOME x. \ p x x x) (SOME x. \ p x x x)))"}), - Rule (Cong, [Rule (Bind, [Rule (Cong, [Rule (Refl, []), Rule (Bind, [Rule (Sko_All, - [Rule (Refl, [])])])])])])); + ((@{term "\ (\x :: nat. p \ (\x :: nat. \x :: nat. q x x))"}, + @{term "\ (\x :: nat. p \ + (\x :: nat. q (SOME x :: nat. \ q x x) (SOME x. \ q x x)))"}), + Node (Cong, [Node (Bind, [Node (Cong, [Node (Refl, []), Node (Bind, [Node (Sko_All, + [Node (Refl, [])])])])])])); reconstruct_proof @{context} proof6 \ @@ -289,7 +289,7 @@ val proof7 = ((@{term "\ \ (\x. p x)"}, @{term "\ \ p (SOME x. p x)"}), - Rule (Cong, [Rule (Cong, [Rule (Sko_Ex, [Rule (Refl, [])])])])); + Node (Cong, [Node (Cong, [Node (Sko_Ex, [Node (Refl, [])])])])); reconstruct_proof @{context} proof7 \ @@ -298,8 +298,8 @@ val proof8 = ((@{term "\ \ (let x = Suc x in x = 0)"}, @{term "\ \ Suc x = 0"}), - Rule (Cong, [Rule (Cong, [Rule (Let [@{term "Suc x"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Cong, [Node (Cong, [Node (Let [@{term "Suc x"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof8 \ @@ -308,7 +308,7 @@ val proof9 = ((@{term "\ (let x = Suc x in x = 0)"}, @{term "\ Suc x = 0"}), - Rule (Cong, [Rule (Let [@{term "Suc x"}], [Rule (Refl, []), Rule (Refl, [])])])); + Node (Cong, [Node (Let [@{term "Suc x"}], [Node (Refl, []), Node (Refl, [])])])); reconstruct_proof @{context} proof9 \ @@ -317,7 +317,7 @@ val proof10 = ((@{term "\x :: nat. p (x + 0)"}, @{term "\x :: nat. p x"}), - Rule (Bind, [Rule (Cong, [Rule (Taut @{thm add_0_right}, [])])])); + Node (Bind, [Node (Cong, [Node (Taut @{thm add_0_right}, [])])])); reconstruct_proof @{context} proof10; \ @@ -326,8 +326,8 @@ val proof11 = ((@{term "\ (let (x, y) = (Suc y, Suc x) in y = 0)"}, @{term "\ Suc x = 0"}), - Rule (Cong, [Rule (Let [@{term "Suc y"}, @{term "Suc x"}], - [Rule (Refl, []), Rule (Refl, []), Rule (Refl, [])])])); + Node (Cong, [Node (Let [@{term "Suc y"}, @{term "Suc x"}], + [Node (Refl, []), Node (Refl, []), Node (Refl, [])])])); reconstruct_proof @{context} proof11 \ @@ -336,10 +336,10 @@ val proof12 = ((@{term "\ (let (x, y) = (Suc y, Suc x); (u, v, w) = (y, x, y) in w = 0)"}, @{term "\ Suc x = 0"}), - Rule (Cong, [Rule (Let [@{term "Suc y"}, @{term "Suc x"}], - [Rule (Refl, []), Rule (Refl, []), - Rule (Let [@{term "Suc x"}, @{term "Suc y"}, @{term "Suc x"}], - [Rule (Refl, []), Rule (Refl, []), Rule (Refl, []), Rule (Refl, [])])])])); + Node (Cong, [Node (Let [@{term "Suc y"}, @{term "Suc x"}], + [Node (Refl, []), Node (Refl, []), + Node (Let [@{term "Suc x"}, @{term "Suc y"}, @{term "Suc x"}], + [Node (Refl, []), Node (Refl, []), Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof12 \ @@ -348,8 +348,8 @@ val proof13 = ((@{term "\ \ (let x = Suc x in x = 0)"}, @{term "\ \ Suc x = 0"}), - Rule (Cong, [Rule (Cong, [Rule (Let [@{term "Suc x"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Cong, [Node (Cong, [Node (Let [@{term "Suc x"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof13 \ @@ -358,8 +358,8 @@ val proof14 = ((@{term "let (x, y) = (f (a :: nat), b :: nat) in x > a"}, @{term "f (a :: nat) > a"}), - Rule (Let [@{term "f (a :: nat) :: nat"}, @{term "b :: nat"}], - [Rule (Cong, [Rule (Refl, [])]), Rule (Refl, []), Rule (Refl, [])])); + Node (Let [@{term "f (a :: nat) :: nat"}, @{term "b :: nat"}], + [Node (Cong, [Node (Refl, [])]), Node (Refl, []), Node (Refl, [])])); reconstruct_proof @{context} proof14 \ @@ -368,9 +368,9 @@ val proof15 = ((@{term "let x = (let y = g (z :: nat) in f (y :: nat)) in x = Suc 0"}, @{term "f (g (z :: nat) :: nat) = Suc 0"}), - Rule (Let [@{term "f (g (z :: nat) :: nat) :: nat"}], - [Rule (Let [@{term "g (z :: nat) :: nat"}], [Rule (Refl, []), Rule (Refl, [])]), - Rule (Refl, [])])); + Node (Let [@{term "f (g (z :: nat) :: nat) :: nat"}], + [Node (Let [@{term "g (z :: nat) :: nat"}], [Node (Refl, []), Node (Refl, [])]), + Node (Refl, [])])); reconstruct_proof @{context} proof15 \ @@ -379,7 +379,7 @@ val proof16 = ((@{term "a > Suc b"}, @{term "a > Suc b"}), - Rule (Trans @{term "a > Suc b"}, [Rule (Refl, []), Rule (Refl, [])])); + Node (Trans @{term "a > Suc b"}, [Node (Refl, []), Node (Refl, [])])); reconstruct_proof @{context} proof16 \ @@ -392,8 +392,8 @@ val proof17 = ((@{term "2 :: nat"}, @{term "Suc (Suc 0) :: nat"}), - Rule (Trans @{term "Suc 1"}, [Rule (Taut @{thm Suc_1[symmetric]}, []), - Rule (Cong, [Rule (Taut @{thm One_nat_def}, [])])])); + Node (Trans @{term "Suc 1"}, [Node (Taut @{thm Suc_1[symmetric]}, []), + Node (Cong, [Node (Taut @{thm One_nat_def}, [])])])); reconstruct_proof @{context} proof17 \ @@ -402,9 +402,9 @@ val proof18 = ((@{term "let x = a in let y = b in Suc x + y"}, @{term "Suc a + b"}), - Rule (Trans @{term "let y = b in Suc a + y"}, - [Rule (Let [@{term "a :: nat"}], [Rule (Refl, []), Rule (Refl, [])]), - Rule (Let [@{term "b :: nat"}], [Rule (Refl, []), Rule (Refl, [])])])); + Node (Trans @{term "let y = b in Suc a + y"}, + [Node (Let [@{term "a :: nat"}], [Node (Refl, []), Node (Refl, [])]), + Node (Let [@{term "b :: nat"}], [Node (Refl, []), Node (Refl, [])])])); reconstruct_proof @{context} proof18 \ @@ -413,8 +413,8 @@ val proof19 = ((@{term "\x. let x = f (x :: nat) :: nat in g x"}, @{term "\x. g (f (x :: nat) :: nat)"}), - Rule (Bind, [Rule (Let [@{term "f :: nat \ nat"} $ Bound 0], - [Rule (Refl, []), Rule (Refl, [])])])); + Node (Bind, [Node (Let [@{term "f :: nat \ nat"} $ Bound 0], + [Node (Refl, []), Node (Refl, [])])])); reconstruct_proof @{context} proof19 \ @@ -423,9 +423,9 @@ val proof20 = ((@{term "\x. let y = Suc 0 in let x = f (x :: nat) :: nat in g x"}, @{term "\x. g (f (x :: nat) :: nat)"}), - Rule (Bind, [Rule (Let [@{term "Suc 0"}], - [Rule (Refl, []), Rule (Let [@{term "f (x :: nat) :: nat"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Bind, [Node (Let [@{term "Suc 0"}], + [Node (Refl, []), Node (Let [@{term "f (x :: nat) :: nat"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof20 \ @@ -434,9 +434,9 @@ val proof21 = ((@{term "\x :: nat. let x = f x :: nat in let y = x in p y"}, @{term "\z :: nat. p (f z :: nat)"}), - Rule (Bind, [Rule (Let [@{term "f (z :: nat) :: nat"}], - [Rule (Refl, []), Rule (Let [@{term "f (z :: nat) :: nat"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Bind, [Node (Let [@{term "f (z :: nat) :: nat"}], + [Node (Refl, []), Node (Let [@{term "f (z :: nat) :: nat"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof21 \ @@ -445,9 +445,9 @@ val proof22 = ((@{term "\x :: nat. let x = f x :: nat in let y = x in p y"}, @{term "\x :: nat. p (f x :: nat)"}), - Rule (Bind, [Rule (Let [@{term "f (x :: nat) :: nat"}], - [Rule (Refl, []), Rule (Let [@{term "f (x :: nat) :: nat"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Bind, [Node (Let [@{term "f (x :: nat) :: nat"}], + [Node (Refl, []), Node (Let [@{term "f (x :: nat) :: nat"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof22 \ @@ -456,9 +456,9 @@ val proof23 = ((@{term "\x :: nat. let (x, a) = (f x :: nat, 0 ::nat) in let y = x in p y"}, @{term "\z :: nat. p (f z :: nat)"}), - Rule (Bind, [Rule (Let [@{term "f (z :: nat) :: nat"}, @{term "0 :: nat"}], - [Rule (Refl, []), Rule (Refl, []), Rule (Let [@{term "f (z :: nat) :: nat"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Bind, [Node (Let [@{term "f (z :: nat) :: nat"}, @{term "0 :: nat"}], + [Node (Refl, []), Node (Refl, []), Node (Let [@{term "f (z :: nat) :: nat"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof23 \ @@ -467,9 +467,9 @@ val proof24 = ((@{term "\x :: nat. let (x, a) = (f x :: nat, 0 ::nat) in let y = x in p y"}, @{term "\x :: nat. p (f x :: nat)"}), - Rule (Bind, [Rule (Let [@{term "f (x :: nat) :: nat"}, @{term "0 :: nat"}], - [Rule (Refl, []), Rule (Refl, []), Rule (Let [@{term "f (x :: nat) :: nat"}], - [Rule (Refl, []), Rule (Refl, [])])])])); + Node (Bind, [Node (Let [@{term "f (x :: nat) :: nat"}, @{term "0 :: nat"}], + [Node (Refl, []), Node (Refl, []), Node (Let [@{term "f (x :: nat) :: nat"}], + [Node (Refl, []), Node (Refl, [])])])])); reconstruct_proof @{context} proof24 \