diff -r c0ab0824ccb5 -r d11249edc2c2 src/HOL/ex/veriT_Preprocessing.thy --- a/src/HOL/ex/veriT_Preprocessing.thy Fri Feb 10 08:27:27 2017 +0100 +++ b/src/HOL/ex/veriT_Preprocessing.thy Fri Feb 10 11:13:19 2017 +0100 @@ -78,7 +78,7 @@ | str_of_rule_name (Let ts) = "Let[" ^ commas (map (Syntax.string_of_term @{context}) ts) ^ "]"; -datatype node = Node of rule_name * node list; +datatype node = N 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), Node (rule_name, prems)) = +fun reconstruct_proof ctxt (lrhs as (_, rhs), N (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)"}), - Node (Sko_Ex, [Node (Refl, [])])); + N (Sko_Ex, [N (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)"}), - Node (Cong, [Node (Sko_All, [Node (Bind, [Node (Refl, [])])])])); + N (Cong, [N (Sko_All, [N (Bind, [N (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)"}), - Node (Bind, [Node (Sko_Ex, [Node (Sko_Ex, [Node (Refl, [])])])])); + N (Bind, [N (Sko_Ex, [N (Sko_Ex, [N (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)"}), - Node (Bind, [Node (Sko_Ex, [Node (Sko_Ex, [Node (Refl, [])])])])); + N (Bind, [N (Sko_Ex, [N (Sko_Ex, [N (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)"}), - Node (Bind, [Node (Bind, [Node (Sko_Ex, [Node (Refl, [])])])])); + N (Bind, [N (Bind, [N (Sko_Ex, [N (Refl, [])])])])); reconstruct_proof @{context} proof4 \ @@ -268,8 +268,7 @@ ((@{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))"}), - Node (Bind, [Node (Cong, [Node (Refl, []), Node (Bind, [Node (Sko_Ex, - [Node (Refl, [])])])])])); + N (Bind, [N (Cong, [N (Refl, []), N (Bind, [N (Sko_Ex, [N (Refl, [])])])])])); reconstruct_proof @{context} proof5 \ @@ -279,8 +278,7 @@ ((@{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, [])])])])])])); + N (Cong, [N (Bind, [N (Cong, [N (Refl, []), N (Bind, [N (Sko_All, [N (Refl, [])])])])])])); reconstruct_proof @{context} proof6 \ @@ -289,7 +287,7 @@ val proof7 = ((@{term "\ \ (\x. p x)"}, @{term "\ \ p (SOME x. p x)"}), - Node (Cong, [Node (Cong, [Node (Sko_Ex, [Node (Refl, [])])])])); + N (Cong, [N (Cong, [N (Sko_Ex, [N (Refl, [])])])])); reconstruct_proof @{context} proof7 \ @@ -298,8 +296,7 @@ val proof8 = ((@{term "\ \ (let x = Suc x in x = 0)"}, @{term "\ \ Suc x = 0"}), - Node (Cong, [Node (Cong, [Node (Let [@{term "Suc x"}], - [Node (Refl, []), Node (Refl, [])])])])); + N (Cong, [N (Cong, [N (Let [@{term "Suc x"}], [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof8 \ @@ -308,7 +305,7 @@ val proof9 = ((@{term "\ (let x = Suc x in x = 0)"}, @{term "\ Suc x = 0"}), - Node (Cong, [Node (Let [@{term "Suc x"}], [Node (Refl, []), Node (Refl, [])])])); + N (Cong, [N (Let [@{term "Suc x"}], [N (Refl, []), N (Refl, [])])])); reconstruct_proof @{context} proof9 \ @@ -317,7 +314,7 @@ val proof10 = ((@{term "\x :: nat. p (x + 0)"}, @{term "\x :: nat. p x"}), - Node (Bind, [Node (Cong, [Node (Taut @{thm add_0_right}, [])])])); + N (Bind, [N (Cong, [N (Taut @{thm add_0_right}, [])])])); reconstruct_proof @{context} proof10; \ @@ -326,8 +323,8 @@ val proof11 = ((@{term "\ (let (x, y) = (Suc y, Suc x) in y = 0)"}, @{term "\ Suc x = 0"}), - Node (Cong, [Node (Let [@{term "Suc y"}, @{term "Suc x"}], - [Node (Refl, []), Node (Refl, []), Node (Refl, [])])])); + N (Cong, [N (Let [@{term "Suc y"}, @{term "Suc x"}], [N (Refl, []), N (Refl, []), + N (Refl, [])])])); reconstruct_proof @{context} proof11 \ @@ -336,10 +333,9 @@ val proof12 = ((@{term "\ (let (x, y) = (Suc y, Suc x); (u, v, w) = (y, x, y) in w = 0)"}, @{term "\ Suc x = 0"}), - 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, [])])])])); + N (Cong, [N (Let [@{term "Suc y"}, @{term "Suc x"}], [N (Refl, []), N (Refl, []), + N (Let [@{term "Suc x"}, @{term "Suc y"}, @{term "Suc x"}], + [N (Refl, []), N (Refl, []), N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof12 \ @@ -348,8 +344,7 @@ val proof13 = ((@{term "\ \ (let x = Suc x in x = 0)"}, @{term "\ \ Suc x = 0"}), - Node (Cong, [Node (Cong, [Node (Let [@{term "Suc x"}], - [Node (Refl, []), Node (Refl, [])])])])); + N (Cong, [N (Cong, [N (Let [@{term "Suc x"}], [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof13 \ @@ -358,8 +353,8 @@ val proof14 = ((@{term "let (x, y) = (f (a :: nat), b :: nat) in x > a"}, @{term "f (a :: nat) > a"}), - Node (Let [@{term "f (a :: nat) :: nat"}, @{term "b :: nat"}], - [Node (Cong, [Node (Refl, [])]), Node (Refl, []), Node (Refl, [])])); + N (Let [@{term "f (a :: nat) :: nat"}, @{term "b :: nat"}], + [N (Cong, [N (Refl, [])]), N (Refl, []), N (Refl, [])])); reconstruct_proof @{context} proof14 \ @@ -368,9 +363,8 @@ 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"}), - Node (Let [@{term "f (g (z :: nat) :: nat) :: nat"}], - [Node (Let [@{term "g (z :: nat) :: nat"}], [Node (Refl, []), Node (Refl, [])]), - Node (Refl, [])])); + N (Let [@{term "f (g (z :: nat) :: nat) :: nat"}], + [N (Let [@{term "g (z :: nat) :: nat"}], [N (Refl, []), N (Refl, [])]), N (Refl, [])])); reconstruct_proof @{context} proof15 \ @@ -379,7 +373,7 @@ val proof16 = ((@{term "a > Suc b"}, @{term "a > Suc b"}), - Node (Trans @{term "a > Suc b"}, [Node (Refl, []), Node (Refl, [])])); + N (Trans @{term "a > Suc b"}, [N (Refl, []), N (Refl, [])])); reconstruct_proof @{context} proof16 \ @@ -392,8 +386,8 @@ val proof17 = ((@{term "2 :: nat"}, @{term "Suc (Suc 0) :: nat"}), - Node (Trans @{term "Suc 1"}, [Node (Taut @{thm Suc_1[symmetric]}, []), - Node (Cong, [Node (Taut @{thm One_nat_def}, [])])])); + N (Trans @{term "Suc 1"}, [N (Taut @{thm Suc_1[symmetric]}, []), N (Cong, + [N (Taut @{thm One_nat_def}, [])])])); reconstruct_proof @{context} proof17 \ @@ -402,9 +396,9 @@ val proof18 = ((@{term "let x = a in let y = b in Suc x + y"}, @{term "Suc a + b"}), - 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, [])])])); + N (Trans @{term "let y = b in Suc a + y"}, + [N (Let [@{term "a :: nat"}], [N (Refl, []), N (Refl, [])]), + N (Let [@{term "b :: nat"}], [N (Refl, []), N (Refl, [])])])); reconstruct_proof @{context} proof18 \ @@ -413,8 +407,8 @@ val proof19 = ((@{term "\x. let x = f (x :: nat) :: nat in g x"}, @{term "\x. g (f (x :: nat) :: nat)"}), - Node (Bind, [Node (Let [@{term "f :: nat \ nat"} $ Bound 0], - [Node (Refl, []), Node (Refl, [])])])); + N (Bind, [N (Let [@{term "f :: nat \ nat"} $ Bound 0], + [N (Refl, []), N (Refl, [])])])); reconstruct_proof @{context} proof19 \ @@ -423,9 +417,8 @@ 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)"}), - Node (Bind, [Node (Let [@{term "Suc 0"}], - [Node (Refl, []), Node (Let [@{term "f (x :: nat) :: nat"}], - [Node (Refl, []), Node (Refl, [])])])])); + N (Bind, [N (Let [@{term "Suc 0"}], [N (Refl, []), N (Let [@{term "f (x :: nat) :: nat"}], + [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof20 \ @@ -434,9 +427,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)"}), - Node (Bind, [Node (Let [@{term "f (z :: nat) :: nat"}], - [Node (Refl, []), Node (Let [@{term "f (z :: nat) :: nat"}], - [Node (Refl, []), Node (Refl, [])])])])); + N (Bind, [N (Let [@{term "f (z :: nat) :: nat"}], + [N (Refl, []), N (Let [@{term "f (z :: nat) :: nat"}], + [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof21 \ @@ -445,9 +438,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)"}), - Node (Bind, [Node (Let [@{term "f (x :: nat) :: nat"}], - [Node (Refl, []), Node (Let [@{term "f (x :: nat) :: nat"}], - [Node (Refl, []), Node (Refl, [])])])])); + N (Bind, [N (Let [@{term "f (x :: nat) :: nat"}], + [N (Refl, []), N (Let [@{term "f (x :: nat) :: nat"}], + [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof22 \ @@ -456,9 +449,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)"}), - 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, [])])])])); + N (Bind, [N (Let [@{term "f (z :: nat) :: nat"}, @{term "0 :: nat"}], + [N (Refl, []), N (Refl, []), N (Let [@{term "f (z :: nat) :: nat"}], + [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof23 \ @@ -467,9 +460,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)"}), - 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, [])])])])); + N (Bind, [N (Let [@{term "f (x :: nat) :: nat"}, @{term "0 :: nat"}], + [N (Refl, []), N (Refl, []), N (Let [@{term "f (x :: nat) :: nat"}], + [N (Refl, []), N (Refl, [])])])])); reconstruct_proof @{context} proof24 \