X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fcoq-contribs%2FLAMBDA-TYPES%2Fbase_types.v;h=f24ef915aafed9d5511fe10390597cb5ede6b78b;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=0b1c870f3c45250a2b15b2b405202df70be62c34;hpb=1b21075e987872a2e3103203b4e67c939e4a9f6a;p=helm.git diff --git a/helm/coq-contribs/LAMBDA-TYPES/base_types.v b/helm/coq-contribs/LAMBDA-TYPES/base_types.v index 0b1c870f3..f24ef915a 100644 --- a/helm/coq-contribs/LAMBDA-TYPES/base_types.v +++ b/helm/coq-contribs/LAMBDA-TYPES/base_types.v @@ -73,6 +73,20 @@ Grammar constr constr10 := [ "EX" ident($v0) ident($v1) "|" constr($c0) "&" constr($c1) "&" constr($c2) ] -> [ (ex3_2 ? ? [$v0;$v1]$c0 [$v0;$v1]$c1 [$v0;$v1]$c2) ]. +(* ex_3 *) + +Inductive ex_3 [A0,A1,A2:Set; P0:A0->A1->A2->Prop] : Prop := + ex_3_intro : (x0:A0; x1:A1; x2:A2)(P0 x0 x1 x2)->(ex_3 A0 A1 A2 P0). + +Hint ex_3 : ltlc := Constructors ex_3. + +Syntactic Definition Ex_3 := ex_3 | 1. + +Grammar constr constr10 := + | ex_3implicit + [ "EX" ident($v0) ident($v1) ident($v2) "|" constr($c0) ] -> + [ (ex_3 ? ? ? [$v0;$v1;$v2]$c0) ]. + (* ex3_3 *) Inductive ex3_3 [A0,A1,A2:Set; P0,P1,P2:A0->A1->A2->Prop] : Prop := @@ -143,6 +157,20 @@ Grammar constr constr10 := [ "EX" ident($v0) ident($v1) ident($v2) ident($v3) ident($v4) "|" constr($c0) "&" constr($c1) "&" constr($c2) "&" constr($c3) ] -> [ (ex4_5 ? ? ? ? ? [$v0;$v1;$v2;$v3;$v4]$c0 [$v0;$v1;$v2;$v3;$v4]$c1 [$v0;$v1;$v2;$v3;$v4]$c2 [$v0;$v1;$v2;$v3;$v4]$c3) ]. +(* ex5_5 *) + +Inductive ex5_5 [A0,A1,A2,A3,A4:Set; P0,P1,P2,P3,P4:A0->A1->A2->A3->A4->Prop] : Prop := + ex5_5_intro : (x0:A0; x1:A1; x2:A2; x3:A3; x4:A4)(P0 x0 x1 x2 x3 x4)->(P1 x0 x1 x2 x3 x4)->(P2 x0 x1 x2 x3 x4)->(P3 x0 x1 x2 x3 x4)->(P4 x0 x1 x2 x3 x4)->(ex5_5 A0 A1 A2 A3 A4 P0 P1 P2 P3 P4). + +Hint ex5_5 : ltlc := Constructors ex5_5. + +Syntactic Definition Ex5_5 := ex5_5 | 1. + +Grammar constr constr10 := + | ex5_5implicit + [ "EX" ident($v0) ident($v1) ident($v2) ident($v3) ident($v4) "|" constr($c0) "&" constr($c1) "&" constr($c2) "&" constr($c3) "&" constr($c4) ] -> + [ (ex5_5 ? ? ? ? ? [$v0;$v1;$v2;$v3;$v4]$c0 [$v0;$v1;$v2;$v3;$v4]$c1 [$v0;$v1;$v2;$v3;$v4]$c2 [$v0;$v1;$v2;$v3;$v4]$c3 [$v0;$v1;$v2;$v3;$v4]$c4) ]. + (* ex6_6 *) Inductive ex6_6 [A0,A1,A2,A3,A4,A5:Set; P0,P1,P2,P3,P4,P5:A0->A1->A2->A3->A4->A5->Prop] : Prop := @@ -157,3 +185,22 @@ Grammar constr constr10 := [ "EX" ident($v0) ident($v1) ident($v2) ident($v3) ident($v4) ident($v5) "|" constr($c0) "&" constr($c1) "&" constr($c2) "&" constr($c3) "&" constr($c4) "&" constr($c5) ] -> [ (ex6_6 ? ? ? ? ? ? [$v0;$v1;$v2;$v3;$v4;$v5]$c0 [$v0;$v1;$v2;$v3;$v4;$v5]$c1 [$v0;$v1;$v2;$v3;$v4;$v5]$c2 [$v0;$v1;$v2;$v3;$v4;$v5]$c3 [$v0;$v1;$v2;$v3;$v4;$v5]$c4 [$v0;$v1;$v2;$v3;$v4;$v5]$c5) ]. +(* ex6_7 *) + +Inductive ex6_7 [A0,A1,A2,A3,A4,A5,A6:Set; P0,P1,P2,P3,P4,P5:A0->A1->A2->A3->A4->A5->A6->Prop] : Prop := + ex6_7_intro : (x0:A0; x1:A1; x2:A2; x3:A3; x4:A4; x5:A5; x6:A6)(P0 x0 x1 x2 x3 x4 x5 x6)->(P1 x0 x1 x2 x3 x4 x5 x6)->(P2 x0 x1 x2 x3 x4 x5 x6)->(P3 x0 x1 x2 x3 x4 x5 x6)->(P4 x0 x1 x2 x3 x4 x5 x6)->(P5 x0 x1 x2 x3 x4 x5 x6)->(ex6_7 A0 A1 A2 A3 A4 A5 A6 P0 P1 P2 P3 P4 P5). + +Hint ex6_7 : ltlc := Constructors ex6_7. + +Syntactic Definition Ex6_7 := ex6_7 | 1. + +Grammar constr constr10 := + | ex6_7implicit + [ "EX" ident($v0) ident($v1) ident($v2) ident($v3) ident($v4) ident($v5) ident($v6) "|" constr($c0) "&" constr($c1) "&" constr($c2) "&" constr($c3) "&" constr($c4) "&" constr($c5) ] -> + [ (ex6_7 ? ? ? ? ? ? ? [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c0 [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c1 [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c2 [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c3 [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c4 [$v0;$v1;$v2;$v3;$v4;$v5;$v6]$c5) ]. + +(* extended Decompose tactic *) + +Tactic Definition XDecompose H := + Decompose [and or ex ex2 or3 or4 ex2_2 ex3_2 ex_3 ex3_3 ex4_3 ex3_4 ex4_4 ex4_5 ex5_5 ex6_6 ex6_7] H; Clear H. +